# HG changeset patch # User wenzelm # Date 1346150749 -7200 # Node ID 12afbf6eb7f9db36d015111e993df99dbf36e8a4 # Parent c04001b3a753bb1a950abac6d20ccb7310c7b4ab more standard document preparation within session context; diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Base.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Base.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,8 @@ +theory Base +imports Pure +begin + +ML_file "../antiquote_setup.ML" +setup Antiquote_Setup.setup + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Document_Preparation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Document_Preparation.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,611 @@ +theory Document_Preparation +imports Base Main +begin + +chapter {* Document preparation \label{ch:document-prep} *} + +text {* Isabelle/Isar provides a simple document preparation system + based on regular {PDF-\LaTeX} technology, with full support for + hyper-links and bookmarks. Thus the results are well suited for WWW + browsing and as printed copies. + + \medskip Isabelle generates {\LaTeX} output while running a + \emph{logic session} (see also \cite{isabelle-sys}). Getting + started with a working configuration for common situations is quite + easy by using the Isabelle @{verbatim mkdir} and @{verbatim make} + tools. First invoke +\begin{ttbox} + isabelle mkdir Foo +\end{ttbox} + to initialize a separate directory for session @{verbatim Foo} (it + is safe to experiment, since @{verbatim "isabelle mkdir"} never + overwrites existing files). Ensure that @{verbatim "Foo/ROOT.ML"} + holds ML commands to load all theories required for this session; + furthermore @{verbatim "Foo/document/root.tex"} should include any + special {\LaTeX} macro packages required for your document (the + default is usually sufficient as a start). + + The session is controlled by a separate @{verbatim IsaMakefile} + (with crude source dependencies by default). This file is located + one level up from the @{verbatim Foo} directory location. Now + invoke +\begin{ttbox} + isabelle make Foo +\end{ttbox} + to run the @{verbatim Foo} session, with browser information and + document preparation enabled. Unless any errors are reported by + Isabelle or {\LaTeX}, the output will appear inside the directory + defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as + reported by the batch job in verbose mode). + + \medskip You may also consider to tune the @{verbatim usedir} + options in @{verbatim IsaMakefile}, for example to switch the output + format between @{verbatim pdf} and @{verbatim dvi}, or activate the + @{verbatim "-D"} option to retain a second copy of the generated + {\LaTeX} sources (for manual inspection or separate runs of + @{executable latex}). + + \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys} + for further details on Isabelle logic sessions and theory + presentation. The Isabelle/HOL tutorial \cite{isabelle-hol-book} + also covers theory presentation to some extent. +*} + + +section {* Markup commands \label{sec:markup} *} + +text {* + \begin{matharray}{rcl} + @{command_def "header"} & : & @{text "toplevel \ toplevel"} \\[0.5ex] + @{command_def "chapter"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "section"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "subsection"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "subsubsection"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "text"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "text_raw"} & : & @{text "local_theory \ local_theory"} \\[0.5ex] + @{command_def "sect"} & : & @{text "proof \ proof"} \\ + @{command_def "subsect"} & : & @{text "proof \ proof"} \\ + @{command_def "subsubsect"} & : & @{text "proof \ proof"} \\ + @{command_def "txt"} & : & @{text "proof \ proof"} \\ + @{command_def "txt_raw"} & : & @{text "proof \ proof"} \\ + \end{matharray} + + Markup commands provide a structured way to insert text into the + document generated from a theory. Each markup command takes a + single @{syntax text} argument, which is passed as argument to a + corresponding {\LaTeX} macro. The default macros provided by + @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according + to the needs of the underlying document and {\LaTeX} styles. + + Note that formal comments (\secref{sec:comments}) are similar to + markup commands, but have a different status within Isabelle/Isar + syntax. + + @{rail " + (@@{command chapter} | @@{command section} | @@{command subsection} | + @@{command subsubsection} | @@{command text}) @{syntax target}? @{syntax text} + ; + (@@{command header} | @@{command text_raw} | @@{command sect} | @@{command subsect} | + @@{command subsubsect} | @@{command txt} | @@{command txt_raw}) @{syntax text} + "} + + \begin{description} + + \item @{command header} provides plain text markup just preceding + the formal beginning of a theory. The corresponding {\LaTeX} macro + is @{verbatim "\\isamarkupheader"}, which acts like @{command + section} by default. + + \item @{command chapter}, @{command section}, @{command subsection}, + and @{command subsubsection} mark chapter and section headings + within the main theory body or local theory targets. The + corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"}, + @{verbatim "\\isamarkupsection"}, @{verbatim + "\\isamarkupsubsection"} etc. + + \item @{command sect}, @{command subsect}, and @{command subsubsect} + mark section headings within proofs. The corresponding {\LaTeX} + macros are @{verbatim "\\isamarkupsect"}, @{verbatim + "\\isamarkupsubsect"} etc. + + \item @{command text} and @{command txt} specify paragraphs of plain + text. This corresponds to a {\LaTeX} environment @{verbatim + "\\begin{isamarkuptext}"} @{text "\"} @{verbatim + "\\end{isamarkuptext}"} etc. + + \item @{command text_raw} and @{command txt_raw} insert {\LaTeX} + source into the output, without additional markup. Thus the full + range of document manipulations becomes available, at the risk of + messing up document output. + + \end{description} + + Except for @{command "text_raw"} and @{command "txt_raw"}, the text + passed to any of the above markup commands may refer to formal + entities via \emph{document antiquotations}, see also + \secref{sec:antiq}. These are interpreted in the present theory or + proof context, or the named @{text "target"}. + + \medskip The proof markup commands closely resemble those for theory + specifications, but have a different formal status and produce + different {\LaTeX} macros. The default definitions coincide for + analogous commands such as @{command section} and @{command sect}. +*} + + +section {* Document Antiquotations \label{sec:antiq} *} + +text {* + \begin{matharray}{rcl} + @{antiquotation_def "theory"} & : & @{text antiquotation} \\ + @{antiquotation_def "thm"} & : & @{text antiquotation} \\ + @{antiquotation_def "lemma"} & : & @{text antiquotation} \\ + @{antiquotation_def "prop"} & : & @{text antiquotation} \\ + @{antiquotation_def "term"} & : & @{text antiquotation} \\ + @{antiquotation_def term_type} & : & @{text antiquotation} \\ + @{antiquotation_def typeof} & : & @{text antiquotation} \\ + @{antiquotation_def const} & : & @{text antiquotation} \\ + @{antiquotation_def abbrev} & : & @{text antiquotation} \\ + @{antiquotation_def typ} & : & @{text antiquotation} \\ + @{antiquotation_def type} & : & @{text antiquotation} \\ + @{antiquotation_def class} & : & @{text antiquotation} \\ + @{antiquotation_def "text"} & : & @{text antiquotation} \\ + @{antiquotation_def goals} & : & @{text antiquotation} \\ + @{antiquotation_def subgoals} & : & @{text antiquotation} \\ + @{antiquotation_def prf} & : & @{text antiquotation} \\ + @{antiquotation_def full_prf} & : & @{text antiquotation} \\ + @{antiquotation_def ML} & : & @{text antiquotation} \\ + @{antiquotation_def ML_op} & : & @{text antiquotation} \\ + @{antiquotation_def ML_type} & : & @{text antiquotation} \\ + @{antiquotation_def ML_struct} & : & @{text antiquotation} \\ + @{antiquotation_def "file"} & : & @{text antiquotation} \\ + \end{matharray} + + The overall content of an Isabelle/Isar theory may alternate between + formal and informal text. The main body consists of formal + specification and proof commands, interspersed with markup commands + (\secref{sec:markup}) or document comments (\secref{sec:comments}). + The argument of markup commands quotes informal text to be printed + in the resulting document, but may again refer to formal entities + via \emph{document antiquotations}. + + For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}'' + within a text block makes + \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document. + + Antiquotations usually spare the author tedious typing of logical + entities in full detail. Even more importantly, some degree of + consistency-checking between the main body of formal text and its + informal explanation is achieved, since terms and types appearing in + antiquotations are checked within the current theory or proof + context. + + %% FIXME less monolithic presentation, move to individual sections!? + @{rail " + '@{' antiquotation '}' + ; + @{syntax_def antiquotation}: + @@{antiquotation theory} options @{syntax name} | + @@{antiquotation thm} options styles @{syntax thmrefs} | + @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? | + @@{antiquotation prop} options styles @{syntax prop} | + @@{antiquotation term} options styles @{syntax term} | + @@{antiquotation (HOL) value} options styles @{syntax term} | + @@{antiquotation term_type} options styles @{syntax term} | + @@{antiquotation typeof} options styles @{syntax term} | + @@{antiquotation const} options @{syntax term} | + @@{antiquotation abbrev} options @{syntax term} | + @@{antiquotation typ} options @{syntax type} | + @@{antiquotation type} options @{syntax name} | + @@{antiquotation class} options @{syntax name} | + @@{antiquotation text} options @{syntax name} + ; + @{syntax antiquotation}: + @@{antiquotation goals} options | + @@{antiquotation subgoals} options | + @@{antiquotation prf} options @{syntax thmrefs} | + @@{antiquotation full_prf} options @{syntax thmrefs} | + @@{antiquotation ML} options @{syntax name} | + @@{antiquotation ML_op} options @{syntax name} | + @@{antiquotation ML_type} options @{syntax name} | + @@{antiquotation ML_struct} options @{syntax name} | + @@{antiquotation \"file\"} options @{syntax name} + ; + options: '[' (option * ',') ']' + ; + option: @{syntax name} | @{syntax name} '=' @{syntax name} + ; + styles: '(' (style + ',') ')' + ; + style: (@{syntax name} +) + "} + + Note that the syntax of antiquotations may \emph{not} include source + comments @{verbatim "(*"}~@{text "\"}~@{verbatim "*)"} nor verbatim + text @{verbatim "{"}@{verbatim "*"}~@{text "\"}~@{verbatim + "*"}@{verbatim "}"}. + + \begin{description} + + \item @{text "@{theory A}"} prints the name @{text "A"}, which is + guaranteed to refer to a valid ancestor theory in the current + context. + + \item @{text "@{thm a\<^sub>1 \ a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \ a\<^sub>n"}. + Full fact expressions are allowed here, including attributes + (\secref{sec:syn-att}). + + \item @{text "@{prop \}"} prints a well-typed proposition @{text + "\"}. + + \item @{text "@{lemma \ by m}"} proves a well-typed proposition + @{text "\"} by method @{text m} and prints the original @{text "\"}. + + \item @{text "@{term t}"} prints a well-typed term @{text "t"}. + + \item @{text "@{value t}"} evaluates a term @{text "t"} and prints + its result, see also @{command_ref (HOL) value}. + + \item @{text "@{term_type t}"} prints a well-typed term @{text "t"} + annotated with its type. + + \item @{text "@{typeof t}"} prints the type of a well-typed term + @{text "t"}. + + \item @{text "@{const c}"} prints a logical or syntactic constant + @{text "c"}. + + \item @{text "@{abbrev c x\<^sub>1 \ x\<^sub>n}"} prints a constant abbreviation + @{text "c x\<^sub>1 \ x\<^sub>n \ rhs"} as defined in the current context. + + \item @{text "@{typ \}"} prints a well-formed type @{text "\"}. + + \item @{text "@{type \}"} prints a (logical or syntactic) type + constructor @{text "\"}. + + \item @{text "@{class c}"} prints a class @{text c}. + + \item @{text "@{text s}"} prints uninterpreted source text @{text + s}. This is particularly useful to print portions of text according + to the Isabelle document style, without demanding well-formedness, + e.g.\ small pieces of terms that should not be parsed or + type-checked yet. + + \item @{text "@{goals}"} prints the current \emph{dynamic} goal + state. This is mainly for support of tactic-emulation scripts + within Isar. Presentation of goal states does not conform to the + idea of human-readable proof documents! + + When explaining proofs in detail it is usually better to spell out + the reasoning via proper Isar proof commands, instead of peeking at + the internal machine configuration. + + \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but + does not print the main goal. + + \item @{text "@{prf a\<^sub>1 \ a\<^sub>n}"} prints the (compact) proof terms + corresponding to the theorems @{text "a\<^sub>1 \ a\<^sub>n"}. Note that this + requires proof terms to be switched on for the current logic + session. + + \item @{text "@{full_prf a\<^sub>1 \ a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \ + a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays + information omitted in the compact proof term, which is denoted by + ``@{text _}'' placeholders there. + + \item @{text "@{ML s}"}, @{text "@{ML_op s}"}, @{text "@{ML_type + s}"}, and @{text "@{ML_struct s}"} check text @{text s} as ML value, + infix operator, type, and structure, respectively. The source is + printed verbatim. + + \item @{text "@{file path}"} checks that @{text "path"} refers to a + file (or directory) and prints it verbatim. + + \end{description} +*} + + +subsection {* Styled antiquotations *} + +text {* The antiquotations @{text thm}, @{text prop} and @{text + term} admit an extra \emph{style} specification to modify the + printed result. A style is specified by a name with a possibly + empty number of arguments; multiple styles can be sequenced with + commas. The following standard styles are available: + + \begin{description} + + \item @{text lhs} extracts the first argument of any application + form with at least two arguments --- typically meta-level or + object-level equality, or any other binary relation. + + \item @{text rhs} is like @{text lhs}, but extracts the second + argument. + + \item @{text "concl"} extracts the conclusion @{text C} from a rule + in Horn-clause normal form @{text "A\<^sub>1 \ \ A\<^sub>n \ C"}. + + \item @{text "prem"} @{text n} extract premise number + @{text "n"} from from a rule in Horn-clause + normal form @{text "A\<^sub>1 \ \ A\<^sub>n \ C"} + + \end{description} +*} + + +subsection {* General options *} + +text {* The following options are available to tune the printed output + of antiquotations. Note that many of these coincide with global ML + flags of the same names. + + \begin{description} + + \item @{antiquotation_option_def show_types}~@{text "= bool"} and + @{antiquotation_option_def show_sorts}~@{text "= bool"} control + printing of explicit type and sort constraints. + + \item @{antiquotation_option_def show_structs}~@{text "= bool"} + controls printing of implicit structures. + + \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"} + controls folding of abbreviations. + + \item @{antiquotation_option_def names_long}~@{text "= bool"} forces + names of types and constants etc.\ to be printed in their fully + qualified internal form. + + \item @{antiquotation_option_def names_short}~@{text "= bool"} + forces names of types and constants etc.\ to be printed unqualified. + Note that internalizing the output again in the current context may + well yield a different result. + + \item @{antiquotation_option_def names_unique}~@{text "= bool"} + determines whether the printed version of qualified names should be + made sufficiently long to avoid overlap with names declared further + back. Set to @{text false} for more concise output. + + \item @{antiquotation_option_def eta_contract}~@{text "= bool"} + prints terms in @{text \}-contracted form. + + \item @{antiquotation_option_def display}~@{text "= bool"} indicates + if the text is to be output as multi-line ``display material'', + rather than a small piece of text without line breaks (which is the + default). + + In this mode the embedded entities are printed in the same style as + the main theory text. + + \item @{antiquotation_option_def break}~@{text "= bool"} controls + line breaks in non-display material. + + \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates + if the output should be enclosed in double quotes. + + \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text + name} to the print mode to be used for presentation. Note that the + standard setup for {\LaTeX} output is already present by default, + including the modes @{text latex} and @{text xsymbols}. + + \item @{antiquotation_option_def margin}~@{text "= nat"} and + @{antiquotation_option_def indent}~@{text "= nat"} change the margin + or indentation for pretty printing of display material. + + \item @{antiquotation_option_def goals_limit}~@{text "= nat"} + determines the maximum number of goals to be printed (for goal-based + antiquotation). + + \item @{antiquotation_option_def source}~@{text "= bool"} prints the + original source text of the antiquotation arguments, rather than its + internal representation. Note that formal checking of + @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still + enabled; use the @{antiquotation "text"} antiquotation for unchecked + output. + + Regular @{text "term"} and @{text "typ"} antiquotations with @{text + "source = false"} involve a full round-trip from the original source + to an internalized logical entity back to a source form, according + to the syntax of the current context. Thus the printed output is + not under direct control of the author, it may even fluctuate a bit + as the underlying theory is changed later on. + + In contrast, @{antiquotation_option source}~@{text "= true"} + admits direct printing of the given source text, with the desirable + well-formedness check in the background, but without modification of + the printed text. + + \end{description} + + For boolean flags, ``@{text "name = true"}'' may be abbreviated as + ``@{text name}''. All of the above flags are disabled by default, + unless changed from ML, say in the @{verbatim "ROOT.ML"} of the + logic session. +*} + + +section {* Markup via command tags \label{sec:tags} *} + +text {* Each Isabelle/Isar command may be decorated by additional + presentation tags, to indicate some modification in the way it is + printed in the document. + + @{rail " + @{syntax_def tags}: ( tag * ) + ; + tag: '%' (@{syntax ident} | @{syntax string}) + "} + + Some tags are pre-declared for certain classes of commands, serving + as default markup if no tags are given in the text: + + \medskip + \begin{tabular}{ll} + @{text "theory"} & theory begin/end \\ + @{text "proof"} & all proof commands \\ + @{text "ML"} & all commands involving ML code \\ + \end{tabular} + + \medskip The Isabelle document preparation system + \cite{isabelle-sys} allows tagged command regions to be presented + specifically, e.g.\ to fold proof texts, or drop parts of the text + completely. + + For example ``@{command "by"}~@{text "%invisible auto"}'' causes + that piece of proof to be treated as @{text invisible} instead of + @{text "proof"} (the default), which may be shown or hidden + depending on the document setup. In contrast, ``@{command + "by"}~@{text "%visible auto"}'' forces this text to be shown + invariably. + + Explicit tag specifications within a proof apply to all subsequent + commands of the same level of nesting. For example, ``@{command + "proof"}~@{text "%visible \"}~@{command "qed"}'' forces the whole + sub-proof to be typeset as @{text visible} (unless some of its parts + are tagged differently). + + \medskip Command tags merely produce certain markup environments for + type-setting. The meaning of these is determined by {\LaTeX} + macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or + by the document author. The Isabelle document preparation tools + also provide some high-level options to specify the meaning of + arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding + parts of the text. Logic sessions may also specify ``document + versions'', where given tags are interpreted in some particular way. + Again see \cite{isabelle-sys} for further details. +*} + + +section {* Railroad diagrams *} + +text {* + \begin{matharray}{rcl} + @{antiquotation_def "rail"} & : & @{text antiquotation} \\ + \end{matharray} + + @{rail "'rail' string"} + + The @{antiquotation rail} antiquotation allows to include syntax + diagrams into Isabelle documents. {\LaTeX} requires the style file + @{"file" "~~/lib/texinputs/pdfsetup.sty"}, which can be used via + @{verbatim "\\usepackage{pdfsetup}"} in @{verbatim "root.tex"}, for + example. + + The rail specification language is quoted here as Isabelle @{syntax + string}; it has its own grammar given below. + + @{rail " + rule? + ';' + ; + rule: ((identifier | @{syntax antiquotation}) ':')? body + ; + body: concatenation + '|' + ; + concatenation: ((atom '?'?) +) (('*' | '+') atom?)? + ; + atom: '(' body? ')' | identifier | + '@'? (string | @{syntax antiquotation}) | + '\\\\\\\\' + "} + + The lexical syntax of @{text "identifier"} coincides with that of + @{syntax ident} in regular Isabelle syntax, but @{text string} uses + single quotes instead of double quotes of the standard @{syntax + string} category, to avoid extra escapes. + + Each @{text rule} defines a formal language (with optional name), + using a notation that is similar to EBNF or regular expressions with + recursion. The meaning and visual appearance of these rail language + elements is illustrated by the following representative examples. + + \begin{itemize} + + \item Empty @{verbatim "()"} + + @{rail "()"} + + \item Nonterminal @{verbatim "A"} + + @{rail "A"} + + \item Nonterminal via Isabelle antiquotation + @{verbatim "@{syntax method}"} + + @{rail "@{syntax method}"} + + \item Terminal @{verbatim "'xyz'"} + + @{rail "'xyz'"} + + \item Terminal in keyword style @{verbatim "@'xyz'"} + + @{rail "@'xyz'"} + + \item Terminal via Isabelle antiquotation + @{verbatim "@@{method rule}"} + + @{rail "@@{method rule}"} + + \item Concatenation @{verbatim "A B C"} + + @{rail "A B C"} + + \item Linebreak @{verbatim "\\\\"} inside + concatenation\footnote{Strictly speaking, this is only a single + backslash, but the enclosing @{syntax string} syntax requires a + second one for escaping.} @{verbatim "A B C \\\\ D E F"} + + @{rail "A B C \\ D E F"} + + \item Variants @{verbatim "A | B | C"} + + @{rail "A | B | C"} + + \item Option @{verbatim "A ?"} + + @{rail "A ?"} + + \item Repetition @{verbatim "A *"} + + @{rail "A *"} + + \item Repetition with separator @{verbatim "A * sep"} + + @{rail "A * sep"} + + \item Strict repetition @{verbatim "A +"} + + @{rail "A +"} + + \item Strict repetition with separator @{verbatim "A + sep"} + + @{rail "A + sep"} + + \end{itemize} +*} + + +section {* Draft presentation *} + +text {* + \begin{matharray}{rcl} + @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + \end{matharray} + + @{rail " + (@@{command display_drafts} | @@{command print_drafts}) (@{syntax name} +) + + "} + + \begin{description} + + \item @{command "display_drafts"}~@{text paths} and @{command + "print_drafts"}~@{text paths} perform simple output of a given list + of raw source files. Only those symbols that do not require + additional {\LaTeX} packages are displayed properly, everything else + is left verbatim. + + \end{description} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/First_Order_Logic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/First_Order_Logic.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,520 @@ + +header {* Example: First-Order Logic *} + +theory %visible First_Order_Logic +imports Base (* FIXME Pure!? *) +begin + +text {* + \noindent In order to commence a new object-logic within + Isabelle/Pure we introduce abstract syntactic categories @{text "i"} + for individuals and @{text "o"} for object-propositions. The latter + is embedded into the language of Pure propositions by means of a + separate judgment. +*} + +typedecl i +typedecl o + +judgment + Trueprop :: "o \ prop" ("_" 5) + +text {* + \noindent Note that the object-logic judgement is implicit in the + syntax: writing @{prop A} produces @{term "Trueprop A"} internally. + From the Pure perspective this means ``@{prop A} is derivable in the + object-logic''. +*} + + +subsection {* Equational reasoning \label{sec:framework-ex-equal} *} + +text {* + Equality is axiomatized as a binary predicate on individuals, with + reflexivity as introduction, and substitution as elimination + principle. Note that the latter is particularly convenient in a + framework like Isabelle, because syntactic congruences are + implicitly produced by unification of @{term "B x"} against + expressions containing occurrences of @{term x}. +*} + +axiomatization + equal :: "i \ i \ o" (infix "=" 50) +where + refl [intro]: "x = x" and + subst [elim]: "x = y \ B x \ B y" + +text {* + \noindent Substitution is very powerful, but also hard to control in + full generality. We derive some common symmetry~/ transitivity + schemes of @{term equal} as particular consequences. +*} + +theorem sym [sym]: + assumes "x = y" + shows "y = x" +proof - + have "x = x" .. + with `x = y` show "y = x" .. +qed + +theorem forw_subst [trans]: + assumes "y = x" and "B x" + shows "B y" +proof - + from `y = x` have "x = y" .. + from this and `B x` show "B y" .. +qed + +theorem back_subst [trans]: + assumes "B x" and "x = y" + shows "B y" +proof - + from `x = y` and `B x` + show "B y" .. +qed + +theorem trans [trans]: + assumes "x = y" and "y = z" + shows "x = z" +proof - + from `y = z` and `x = y` + show "x = z" .. +qed + + +subsection {* Basic group theory *} + +text {* + As an example for equational reasoning we consider some bits of + group theory. The subsequent locale definition postulates group + operations and axioms; we also derive some consequences of this + specification. +*} + +locale group = + fixes prod :: "i \ i \ i" (infix "\" 70) + and inv :: "i \ i" ("(_\)" [1000] 999) + and unit :: i ("1") + assumes assoc: "(x \ y) \ z = x \ (y \ z)" + and left_unit: "1 \ x = x" + and left_inv: "x\ \ x = 1" +begin + +theorem right_inv: "x \ x\ = 1" +proof - + have "x \ x\ = 1 \ (x \ x\)" by (rule left_unit [symmetric]) + also have "\ = (1 \ x) \ x\" by (rule assoc [symmetric]) + also have "1 = (x\)\ \ x\" by (rule left_inv [symmetric]) + also have "\ \ x = (x\)\ \ (x\ \ x)" by (rule assoc) + also have "x\ \ x = 1" by (rule left_inv) + also have "((x\)\ \ \) \ x\ = (x\)\ \ (1 \ x\)" by (rule assoc) + also have "1 \ x\ = x\" by (rule left_unit) + also have "(x\)\ \ \ = 1" by (rule left_inv) + finally show "x \ x\ = 1" . +qed + +theorem right_unit: "x \ 1 = x" +proof - + have "1 = x\ \ x" by (rule left_inv [symmetric]) + also have "x \ \ = (x \ x\) \ x" by (rule assoc [symmetric]) + also have "x \ x\ = 1" by (rule right_inv) + also have "\ \ x = x" by (rule left_unit) + finally show "x \ 1 = x" . +qed + +text {* + \noindent Reasoning from basic axioms is often tedious. Our proofs + work by producing various instances of the given rules (potentially + the symmetric form) using the pattern ``@{command have}~@{text + eq}~@{command "by"}~@{text "(rule r)"}'' and composing the chain of + results via @{command also}/@{command finally}. These steps may + involve any of the transitivity rules declared in + \secref{sec:framework-ex-equal}, namely @{thm trans} in combining + the first two results in @{thm right_inv} and in the final steps of + both proofs, @{thm forw_subst} in the first combination of @{thm + right_unit}, and @{thm back_subst} in all other calculational steps. + + Occasional substitutions in calculations are adequate, but should + not be over-emphasized. The other extreme is to compose a chain by + plain transitivity only, with replacements occurring always in + topmost position. For example: +*} + +(*<*) +theorem "\A. PROP A \ PROP A" +proof - + assume [symmetric, defn]: "\x y. (x \ y) \ Trueprop (x = y)" +(*>*) + have "x \ 1 = x \ (x\ \ x)" unfolding left_inv .. + also have "\ = (x \ x\) \ x" unfolding assoc .. + also have "\ = 1 \ x" unfolding right_inv .. + also have "\ = x" unfolding left_unit .. + finally have "x \ 1 = x" . +(*<*) +qed +(*>*) + +text {* + \noindent Here we have re-used the built-in mechanism for unfolding + definitions in order to normalize each equational problem. A more + realistic object-logic would include proper setup for the Simplifier + (\secref{sec:simplifier}), the main automated tool for equational + reasoning in Isabelle. Then ``@{command unfolding}~@{thm + left_inv}~@{command ".."}'' would become ``@{command "by"}~@{text + "(simp only: left_inv)"}'' etc. +*} + +end + + +subsection {* Propositional logic \label{sec:framework-ex-prop} *} + +text {* + We axiomatize basic connectives of propositional logic: implication, + disjunction, and conjunction. The associated rules are modeled + after Gentzen's system of Natural Deduction \cite{Gentzen:1935}. +*} + +axiomatization + imp :: "o \ o \ o" (infixr "\" 25) where + impI [intro]: "(A \ B) \ A \ B" and + impD [dest]: "(A \ B) \ A \ B" + +axiomatization + disj :: "o \ o \ o" (infixr "\" 30) where + disjI\<^isub>1 [intro]: "A \ A \ B" and + disjI\<^isub>2 [intro]: "B \ A \ B" and + disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" + +axiomatization + conj :: "o \ o \ o" (infixr "\" 35) where + conjI [intro]: "A \ B \ A \ B" and + conjD\<^isub>1: "A \ B \ A" and + conjD\<^isub>2: "A \ B \ B" + +text {* + \noindent The conjunctive destructions have the disadvantage that + decomposing @{prop "A \ B"} involves an immediate decision which + component should be projected. The more convenient simultaneous + elimination @{prop "A \ B \ (A \ B \ C) \ C"} can be derived as + follows: +*} + +theorem conjE [elim]: + assumes "A \ B" + obtains A and B +proof + from `A \ B` show A by (rule conjD\<^isub>1) + from `A \ B` show B by (rule conjD\<^isub>2) +qed + +text {* + \noindent Here is an example of swapping conjuncts with a single + intermediate elimination step: +*} + +(*<*) +lemma "\A. PROP A \ PROP A" +proof - +(*>*) + assume "A \ B" + then obtain B and A .. + then have "B \ A" .. +(*<*) +qed +(*>*) + +text {* + \noindent Note that the analogous elimination rule for disjunction + ``@{text "\ A \ B \ A \ B"}'' coincides with + the original axiomatization of @{thm disjE}. + + \medskip We continue propositional logic by introducing absurdity + with its characteristic elimination. Plain truth may then be + defined as a proposition that is trivially true. +*} + +axiomatization + false :: o ("\") where + falseE [elim]: "\ \ A" + +definition + true :: o ("\") where + "\ \ \ \ \" + +theorem trueI [intro]: \ + unfolding true_def .. + +text {* + \medskip\noindent Now negation represents an implication towards + absurdity: +*} + +definition + not :: "o \ o" ("\ _" [40] 40) where + "\ A \ A \ \" + +theorem notI [intro]: + assumes "A \ \" + shows "\ A" +unfolding not_def +proof + assume A + then show \ by (rule `A \ \`) +qed + +theorem notE [elim]: + assumes "\ A" and A + shows B +proof - + from `\ A` have "A \ \" unfolding not_def . + from `A \ \` and `A` have \ .. + then show B .. +qed + + +subsection {* Classical logic *} + +text {* + Subsequently we state the principle of classical contradiction as a + local assumption. Thus we refrain from forcing the object-logic + into the classical perspective. Within that context, we may derive + well-known consequences of the classical principle. +*} + +locale classical = + assumes classical: "(\ C \ C) \ C" +begin + +theorem double_negation: + assumes "\ \ C" + shows C +proof (rule classical) + assume "\ C" + with `\ \ C` show C .. +qed + +theorem tertium_non_datur: "C \ \ C" +proof (rule double_negation) + show "\ \ (C \ \ C)" + proof + assume "\ (C \ \ C)" + have "\ C" + proof + assume C then have "C \ \ C" .. + with `\ (C \ \ C)` show \ .. + qed + then have "C \ \ C" .. + with `\ (C \ \ C)` show \ .. + qed +qed + +text {* + \noindent These examples illustrate both classical reasoning and + non-trivial propositional proofs in general. All three rules + characterize classical logic independently, but the original rule is + already the most convenient to use, because it leaves the conclusion + unchanged. Note that @{prop "(\ C \ C) \ C"} fits again into our + format for eliminations, despite the additional twist that the + context refers to the main conclusion. So we may write @{thm + classical} as the Isar statement ``@{text "\ \ thesis"}''. + This also explains nicely how classical reasoning really works: + whatever the main @{text thesis} might be, we may always assume its + negation! +*} + +end + + +subsection {* Quantifiers \label{sec:framework-ex-quant} *} + +text {* + Representing quantifiers is easy, thanks to the higher-order nature + of the underlying framework. According to the well-known technique + introduced by Church \cite{church40}, quantifiers are operators on + predicates, which are syntactically represented as @{text "\"}-terms + of type @{typ "i \ o"}. Binder notation turns @{text "All (\x. B + x)"} into @{text "\x. B x"} etc. +*} + +axiomatization + All :: "(i \ o) \ o" (binder "\" 10) where + allI [intro]: "(\x. B x) \ \x. B x" and + allD [dest]: "(\x. B x) \ B a" + +axiomatization + Ex :: "(i \ o) \ o" (binder "\" 10) where + exI [intro]: "B a \ (\x. B x)" and + exE [elim]: "(\x. B x) \ (\x. B x \ C) \ C" + +text {* + \noindent The statement of @{thm exE} corresponds to ``@{text + "\ \x. B x \ x \ B x"}'' in Isar. In the + subsequent example we illustrate quantifier reasoning involving all + four rules: +*} + +theorem + assumes "\x. \y. R x y" + shows "\y. \x. R x y" +proof -- {* @{text "\"} introduction *} + obtain x where "\y. R x y" using `\x. \y. R x y` .. -- {* @{text "\"} elimination *} + fix y have "R x y" using `\y. R x y` .. -- {* @{text "\"} destruction *} + then show "\x. R x y" .. -- {* @{text "\"} introduction *} +qed + + +subsection {* Canonical reasoning patterns *} + +text {* + The main rules of first-order predicate logic from + \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} + can now be summarized as follows, using the native Isar statement + format of \secref{sec:framework-stmt}. + + \medskip + \begin{tabular}{l} + @{text "impI: \ A \ B \ A \ B"} \\ + @{text "impD: \ A \ B \ A \ B"} \\[1ex] + + @{text "disjI\<^isub>1: \ A \ A \ B"} \\ + @{text "disjI\<^isub>2: \ B \ A \ B"} \\ + @{text "disjE: \ A \ B \ A \ B"} \\[1ex] + + @{text "conjI: \ A \ B \ A \ B"} \\ + @{text "conjE: \ A \ B \ A \ B"} \\[1ex] + + @{text "falseE: \ \ \ A"} \\ + @{text "trueI: \ \"} \\[1ex] + + @{text "notI: \ A \ \ \ \ A"} \\ + @{text "notE: \ \ A \ A \ B"} \\[1ex] + + @{text "allI: \ \x. B x \ \x. B x"} \\ + @{text "allE: \ \x. B x \ B a"} \\[1ex] + + @{text "exI: \ B a \ \x. B x"} \\ + @{text "exE: \ \x. B x \ a \ B a"} + \end{tabular} + \medskip + + \noindent This essentially provides a declarative reading of Pure + rules as Isar reasoning patterns: the rule statements tells how a + canonical proof outline shall look like. Since the above rules have + already been declared as @{attribute (Pure) intro}, @{attribute + (Pure) elim}, @{attribute (Pure) dest} --- each according to its + particular shape --- we can immediately write Isar proof texts as + follows: +*} + +(*<*) +theorem "\A. PROP A \ PROP A" +proof - +(*>*) + + txt_raw {*\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" + proof + assume A + show B sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" and A sorry %noproof + then have B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have A sorry %noproof + then have "A \ B" .. + + have B sorry %noproof + then have "A \ B" .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" sorry %noproof + then have C + proof + assume A + then show C sorry %noproof + next + assume B + then show C sorry %noproof + qed + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have A and B sorry %noproof + then have "A \ B" .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "A \ B" sorry %noproof + then obtain A and B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\" sorry %noproof + then have A .. + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\" .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\ A" + proof + assume A + then show "\" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\ A" and A sorry %noproof + then have B .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" + proof + fix x + show "B x" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" sorry %noproof + then have "B a" .. + + txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" + proof + show "B a" sorry %noproof + qed + + txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) + + have "\x. B x" sorry %noproof + then obtain a where "B a" .. + + txt_raw {*\end{minipage}*} + +(*<*) +qed +(*>*) + +text {* + \bigskip\noindent Of course, these proofs are merely examples. As + sketched in \secref{sec:framework-subproof}, there is a fair amount + of flexibility in expressing Pure deductions in Isar. Here the user + is asked to express himself adequately, aiming at proof texts of + literary quality. +*} + +end %visible diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Framework.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Framework.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1016 @@ +theory Framework +imports Base Main +begin + +chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *} + +text {* + Isabelle/Isar + \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift} + is intended as a generic framework for developing formal + mathematical documents with full proof checking. Definitions and + proofs are organized as theories. An assembly of theory sources may + be presented as a printed document; see also + \chref{ch:document-prep}. + + The main objective of Isar is the design of a human-readable + structured proof language, which is called the ``primary proof + format'' in Isar terminology. Such a primary proof language is + somewhere in the middle between the extremes of primitive proof + objects and actual natural language. In this respect, Isar is a bit + more formalistic than Mizar + \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar}, + using logical symbols for certain reasoning schemes where Mizar + would prefer English words; see \cite{Wenzel-Wiedijk:2002} for + further comparisons of these systems. + + So Isar challenges the traditional way of recording informal proofs + in mathematical prose, as well as the common tendency to see fully + formal proofs directly as objects of some logical calculus (e.g.\ + @{text "\"}-terms in a version of type theory). In fact, Isar is + better understood as an interpreter of a simple block-structured + language for describing the data flow of local facts and goals, + interspersed with occasional invocations of proof methods. + Everything is reduced to logical inferences internally, but these + steps are somewhat marginal compared to the overall bookkeeping of + the interpretation process. Thanks to careful design of the syntax + and semantics of Isar language elements, a formal record of Isar + instructions may later appear as an intelligible text to the + attentive reader. + + The Isar proof language has emerged from careful analysis of some + inherent virtues of the existing logical framework of Isabelle/Pure + \cite{paulson-found,paulson700}, notably composition of higher-order + natural deduction rules, which is a generalization of Gentzen's + original calculus \cite{Gentzen:1935}. The approach of generic + inference systems in Pure is continued by Isar towards actual proof + texts. + + Concrete applications require another intermediate layer: an + object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed + set-theory) is being used most of the time; Isabelle/ZF + \cite{isabelle-ZF} is less extensively developed, although it would + probably fit better for classical mathematics. + + \medskip In order to illustrate natural deduction in Isar, we shall + refer to the background theory and library of Isabelle/HOL. This + includes common notions of predicate logic, naive set-theory etc.\ + using fairly standard mathematical notation. From the perspective + of generic natural deduction there is nothing special about the + logical connectives of HOL (@{text "\"}, @{text "\"}, @{text "\"}, + @{text "\"}, etc.), only the resulting reasoning principles are + relevant to the user. There are similar rules available for + set-theory operators (@{text "\"}, @{text "\"}, @{text "\"}, @{text + "\"}, etc.), or any other theory developed in the library (lattice + theory, topology etc.). + + Subsequently we briefly review fragments of Isar proof texts + corresponding directly to such general deduction schemes. The + examples shall refer to set-theory, to minimize the danger of + understanding connectives of predicate logic as something special. + + \medskip The following deduction performs @{text "\"}-introduction, + working forwards from assumptions towards the conclusion. We give + both the Isar text, and depict the primitive rule involved, as + determined by unification of the problem against rules that are + declared in the library context. +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +notepad +begin +(*>*) + assume "x \ A" and "x \ B" + then have "x \ A \ B" .. +(*<*) +end +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "x \ A \ B"}}{@{prop "x \ A"} & @{prop "x \ B"}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Note that @{command assume} augments the proof + context, @{command then} indicates that the current fact shall be + used in the next step, and @{command have} states an intermediate + goal. The two dots ``@{command ".."}'' refer to a complete proof of + this claim, using the indicated facts and a canonical rule from the + context. We could have been more explicit here by spelling out the + final proof step via the @{command "by"} command: +*} + +(*<*) +notepad +begin +(*>*) + assume "x \ A" and "x \ B" + then have "x \ A \ B" by (rule IntI) +(*<*) +end +(*>*) + +text {* + \noindent The format of the @{text "\"}-introduction rule represents + the most basic inference, which proceeds from given premises to a + conclusion, without any nested proof context involved. + + The next example performs backwards introduction on @{term "\\"}, + the intersection of all sets within a given set. This requires a + nested proof of set membership within a local context, where @{term + A} is an arbitrary-but-fixed member of the collection: +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +notepad +begin +(*>*) + have "x \ \\" + proof + fix A + assume "A \ \" + show "x \ A" sorry %noproof + qed +(*<*) +end +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "x \ \\"}}{\infer*{@{prop "x \ A"}}{@{text "[A][A \ \]"}}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent This Isar reasoning pattern again refers to the + primitive rule depicted above. The system determines it in the + ``@{command proof}'' step, which could have been spelt out more + explicitly as ``@{command proof}~@{text "(rule InterI)"}''. Note + that the rule involves both a local parameter @{term "A"} and an + assumption @{prop "A \ \"} in the nested reasoning. This kind of + compound rule typically demands a genuine sub-proof in Isar, working + backwards rather than forwards as seen before. In the proof body we + encounter the @{command fix}-@{command assume}-@{command show} + outline of nested sub-proofs that is typical for Isar. The final + @{command show} is like @{command have} followed by an additional + refinement of the enclosing claim, using the rule derived from the + proof body. + + \medskip The next example involves @{term "\\"}, which can be + characterized as the set of all @{term "x"} such that @{prop "\A. x + \ A \ A \ \"}. The elimination rule for @{prop "x \ \\"} does + not mention @{text "\"} and @{text "\"} at all, but admits to obtain + directly a local @{term "A"} such that @{prop "x \ A"} and @{prop "A + \ \"} hold. This corresponds to the following Isar proof and + inference rule, respectively: +*} + +text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} + +(*<*) +notepad +begin +(*>*) + assume "x \ \\" + then have C + proof + fix A + assume "x \ A" and "A \ \" + show C sorry %noproof + qed +(*<*) +end +(*>*) + +text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} + +text {* + \infer{@{prop "C"}}{@{prop "x \ \\"} & \infer*{@{prop "C"}~}{@{text "[A][x \ A, A \ \]"}}} +*} + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Although the Isar proof follows the natural + deduction rule closely, the text reads not as natural as + anticipated. There is a double occurrence of an arbitrary + conclusion @{prop "C"}, which represents the final result, but is + irrelevant for now. This issue arises for any elimination rule + involving local parameters. Isar provides the derived language + element @{command obtain}, which is able to perform the same + elimination proof more conveniently: +*} + +(*<*) +notepad +begin +(*>*) + assume "x \ \\" + then obtain A where "x \ A" and "A \ \" .. +(*<*) +end +(*>*) + +text {* + \noindent Here we avoid to mention the final conclusion @{prop "C"} + and return to plain forward reasoning. The rule involved in the + ``@{command ".."}'' proof is the same as before. +*} + + +section {* The Pure framework \label{sec:framework-pure} *} + +text {* + The Pure logic \cite{paulson-found,paulson700} is an intuitionistic + fragment of higher-order logic \cite{church40}. In type-theoretic + parlance, there are three levels of @{text "\"}-calculus with + corresponding arrows @{text "\"}/@{text "\"}/@{text "\"}: + + \medskip + \begin{tabular}{ll} + @{text "\ \ \"} & syntactic function space (terms depending on terms) \\ + @{text "\x. B(x)"} & universal quantification (proofs depending on terms) \\ + @{text "A \ B"} & implication (proofs depending on proofs) \\ + \end{tabular} + \medskip + + \noindent Here only the types of syntactic terms, and the + propositions of proof terms have been shown. The @{text + "\"}-structure of proofs can be recorded as an optional feature of + the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but + the formal system can never depend on them due to \emph{proof + irrelevance}. + + On top of this most primitive layer of proofs, Pure implements a + generic calculus for nested natural deduction rules, similar to + \cite{Schroeder-Heister:1984}. Here object-logic inferences are + internalized as formulae over @{text "\"} and @{text "\"}. + Combining such rule statements may involve higher-order unification + \cite{paulson-natural}. +*} + + +subsection {* Primitive inferences *} + +text {* + Term syntax provides explicit notation for abstraction @{text "\x :: + \. b(x)"} and application @{text "b a"}, while types are usually + implicit thanks to type-inference; terms of type @{text "prop"} are + called propositions. Logical statements are composed via @{text "\x + :: \. B(x)"} and @{text "A \ B"}. Primitive reasoning operates on + judgments of the form @{text "\ \ \"}, with standard introduction + and elimination rules for @{text "\"} and @{text "\"} that refer to + fixed parameters @{text "x\<^isub>1, \, x\<^isub>m"} and hypotheses + @{text "A\<^isub>1, \, A\<^isub>n"} from the context @{text "\"}; + the corresponding proof terms are left implicit. The subsequent + inference rules define @{text "\ \ \"} inductively, relative to a + collection of axioms: + + \[ + \infer{@{text "\ A"}}{(@{text "A"} \text{~axiom})} + \qquad + \infer{@{text "A \ A"}}{} + \] + + \[ + \infer{@{text "\ \ \x. B(x)"}}{@{text "\ \ B(x)"} & @{text "x \ \"}} + \qquad + \infer{@{text "\ \ B(a)"}}{@{text "\ \ \x. B(x)"}} + \] + + \[ + \infer{@{text "\ - A \ A \ B"}}{@{text "\ \ B"}} + \qquad + \infer{@{text "\\<^sub>1 \ \\<^sub>2 \ B"}}{@{text "\\<^sub>1 \ A \ B"} & @{text "\\<^sub>2 \ A"}} + \] + + Furthermore, Pure provides a built-in equality @{text "\ :: \ \ \ \ + prop"} with axioms for reflexivity, substitution, extensionality, + and @{text "\\\"}-conversion on @{text "\"}-terms. + + \medskip An object-logic introduces another layer on top of Pure, + e.g.\ with types @{text "i"} for individuals and @{text "o"} for + propositions, term constants @{text "Trueprop :: o \ prop"} as + (implicit) derivability judgment and connectives like @{text "\ :: o + \ o \ o"} or @{text "\ :: (i \ o) \ o"}, and axioms for object-level + rules such as @{text "conjI: A \ B \ A \ B"} or @{text "allI: (\x. B + x) \ \x. B x"}. Derived object rules are represented as theorems of + Pure. After the initial object-logic setup, further axiomatizations + are usually avoided; plain definitions and derived principles are + used exclusively. +*} + + +subsection {* Reasoning with rules \label{sec:framework-resolution} *} + +text {* + Primitive inferences mostly serve foundational purposes. The main + reasoning mechanisms of Pure operate on nested natural deduction + rules expressed as formulae, using @{text "\"} to bind local + parameters and @{text "\"} to express entailment. Multiple + parameters and premises are represented by repeating these + connectives in a right-associative manner. + + Since @{text "\"} and @{text "\"} commute thanks to the theorem + @{prop "(A \ (\x. B x)) \ (\x. A \ B x)"}, we may assume w.l.o.g.\ + that rule statements always observe the normal form where + quantifiers are pulled in front of implications at each level of + nesting. This means that any Pure proposition may be presented as a + \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the + form @{text "\x\<^isub>1 \ x\<^isub>m. H\<^isub>1 \ \ H\<^isub>n \ + A"} for @{text "m, n \ 0"}, and @{text "A"} atomic, and @{text + "H\<^isub>1, \, H\<^isub>n"} being recursively of the same format. + Following the convention that outermost quantifiers are implicit, + Horn clauses @{text "A\<^isub>1 \ \ A\<^isub>n \ A"} are a special + case of this. + + For example, @{text "\"}-introduction rule encountered before is + represented as a Pure theorem as follows: + \[ + @{text "IntI:"}~@{prop "x \ A \ x \ B \ x \ A \ B"} + \] + + \noindent This is a plain Horn clause, since no further nesting on + the left is involved. The general @{text "\"}-introduction + corresponds to a Hereditary Harrop Formula with one additional level + of nesting: + \[ + @{text "InterI:"}~@{prop "(\A. A \ \ \ x \ A) \ x \ \\"} + \] + + \medskip Goals are also represented as rules: @{text "A\<^isub>1 \ + \ A\<^isub>n \ C"} states that the sub-goals @{text "A\<^isub>1, \, + A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the + goal is finished. To allow @{text "C"} being a rule statement + itself, we introduce the protective marker @{text "# :: prop \ + prop"}, which is defined as identity and hidden from the user. We + initialize and finish goal states as follows: + + \[ + \begin{array}{c@ {\qquad}c} + \infer[(@{inference_def init})]{@{text "C \ #C"}}{} & + \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}} + \end{array} + \] + + \noindent Goal states are refined in intermediate proof steps until + a finished form is achieved. Here the two main reasoning principles + are @{inference resolution}, for back-chaining a rule against a + sub-goal (replacing it by zero or more sub-goals), and @{inference + assumption}, for solving a sub-goal (finding a short-circuit with + local assumptions). Below @{text "\<^vec>x"} stands for @{text + "x\<^isub>1, \, x\<^isub>n"} (@{text "n \ 0"}). + + \[ + \infer[(@{inference_def resolution})] + {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (\<^vec>a \<^vec>x))\ \ C\"}} + {\begin{tabular}{rl} + @{text "rule:"} & + @{text "\<^vec>A \<^vec>a \ B \<^vec>a"} \\ + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ + @{text "goal unifier:"} & + @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ + \end{tabular}} + \] + + \medskip + + \[ + \infer[(@{inference_def assumption})]{@{text "C\"}} + {\begin{tabular}{rl} + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C"} \\ + @{text "assm unifier:"} & @{text "A\ = H\<^sub>i\"}~~\text{(for some~@{text "H\<^sub>i"})} \\ + \end{tabular}} + \] + + The following trace illustrates goal-oriented reasoning in + Isabelle/Pure: + + {\footnotesize + \medskip + \begin{tabular}{r@ {\quad}l} + @{text "(A \ B \ B \ A) \ #(A \ B \ B \ A)"} & @{text "(init)"} \\ + @{text "(A \ B \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution B \ A \ B \ A)"} \\ + @{text "(A \ B \ A \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution A \ B \ B)"} \\ + @{text "(A \ B \ A) \ #\"} & @{text "(assumption)"} \\ + @{text "(A \ B \ A \ B) \ #\"} & @{text "(resolution A \ B \ A)"} \\ + @{text "#\"} & @{text "(assumption)"} \\ + @{text "A \ B \ B \ A"} & @{text "(finish)"} \\ + \end{tabular} + \medskip + } + + Compositions of @{inference assumption} after @{inference + resolution} occurs quite often, typically in elimination steps. + Traditional Isabelle tactics accommodate this by a combined + @{inference_def elim_resolution} principle. In contrast, Isar uses + a slightly more refined combination, where the assumptions to be + closed are marked explicitly, using again the protective marker + @{text "#"}: + + \[ + \infer[(@{inference refinement})] + {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>G' (\<^vec>a \<^vec>x))\ \ C\"}} + {\begin{tabular}{rl} + @{text "sub\proof:"} & + @{text "\<^vec>G \<^vec>a \ B \<^vec>a"} \\ + @{text "goal:"} & + @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ + @{text "goal unifier:"} & + @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ + @{text "assm unifiers:"} & + @{text "(\\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\ = #H\<^sub>i\"} \\ + & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\ + \end{tabular}} + \] + + \noindent Here the @{text "sub\proof"} rule stems from the + main @{command fix}-@{command assume}-@{command show} outline of + Isar (cf.\ \secref{sec:framework-subproof}): each assumption + indicated in the text results in a marked premise @{text "G"} above. + The marking enforces resolution against one of the sub-goal's + premises. Consequently, @{command fix}-@{command assume}-@{command + show} enables to fit the result of a sub-proof quite robustly into a + pending sub-goal, while maintaining a good measure of flexibility. +*} + + +section {* The Isar proof language \label{sec:framework-isar} *} + +text {* + Structured proofs are presented as high-level expressions for + composing entities of Pure (propositions, facts, and goals). The + Isar proof language allows to organize reasoning within the + underlying rule calculus of Pure, but Isar is not another logical + calculus! + + Isar is an exercise in sound minimalism. Approximately half of the + language is introduced as primitive, the rest defined as derived + concepts. The following grammar describes the core language + (category @{text "proof"}), which is embedded into theory + specification elements such as @{command theorem}; see also + \secref{sec:framework-stmt} for the separate category @{text + "statement"}. + + \medskip + \begin{tabular}{rcl} + @{text "theory\stmt"} & = & @{command "theorem"}~@{text "statement proof |"}~~@{command "definition"}~@{text "\ | \"} \\[1ex] + + @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex] + + @{text prfx} & = & @{command "using"}~@{text "facts"} \\ + & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ + + @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ + & @{text "|"} & @{command "next"} \\ + & @{text "|"} & @{command "note"}~@{text "name = facts"} \\ + & @{text "|"} & @{command "let"}~@{text "term = term"} \\ + & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ + & @{text "|"} & @{command assume}~@{text "\inference\ name: props"} \\ + & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ + @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ + & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ + \end{tabular} + + \medskip Simultaneous propositions or facts may be separated by the + @{keyword "and"} keyword. + + \medskip The syntax for terms and propositions is inherited from + Pure (and the object-logic). A @{text "pattern"} is a @{text + "term"} with schematic variables, to be bound by higher-order + matching. + + \medskip Facts may be referenced by name or proposition. For + example, the result of ``@{command have}~@{text "a: A \proof\"}'' + becomes available both as @{text "a"} and + \isacharbackquoteopen@{text "A"}\isacharbackquoteclose. Moreover, + fact expressions may involve attributes that modify either the + theorem or the background context. For example, the expression + ``@{text "a [OF b]"}'' refers to the composition of two facts + according to the @{inference resolution} inference of + \secref{sec:framework-resolution}, while ``@{text "a [intro]"}'' + declares a fact as introduction rule in the context. + + The special fact called ``@{fact this}'' always refers to the last + result, as produced by @{command note}, @{command assume}, @{command + have}, or @{command show}. Since @{command note} occurs + frequently together with @{command then} we provide some + abbreviations: + + \medskip + \begin{tabular}{rcl} + @{command from}~@{text a} & @{text "\"} & @{command note}~@{text a}~@{command then} \\ + @{command with}~@{text a} & @{text "\"} & @{command from}~@{text "a \ this"} \\ + \end{tabular} + \medskip + + The @{text "method"} category is essentially a parameter and may be + populated later. Methods use the facts indicated by @{command + "then"} or @{command using}, and then operate on the goal state. + Some basic methods are predefined: ``@{method "-"}'' leaves the goal + unchanged, ``@{method this}'' applies the facts as rules to the + goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the + result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}'' + refer to @{inference resolution} of + \secref{sec:framework-resolution}). The secondary arguments to + ``@{method (Pure) rule}'' may be specified explicitly as in ``@{text "(rule + a)"}'', or picked from the context. In the latter case, the system + first tries rules declared as @{attribute (Pure) elim} or + @{attribute (Pure) dest}, followed by those declared as @{attribute + (Pure) intro}. + + The default method for @{command proof} is ``@{method (Pure) rule}'' + (arguments picked from the context), for @{command qed} it is + ``@{method "-"}''. Further abbreviations for terminal proof steps + are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for + ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text + "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command + "by"}~@{method (Pure) rule}, and ``@{command "."}'' for ``@{command + "by"}~@{method this}''. The @{command unfolding} element operates + directly on the current facts and goal by applying equalities. + + \medskip Block structure can be indicated explicitly by ``@{command + "{"}~@{text "\"}~@{command "}"}'', although the body of a sub-proof + already involves implicit nesting. In any case, @{command next} + jumps into the next section of a block, i.e.\ it acts like closing + an implicit block scope and opening another one; there is no direct + correspondence to subgoals here. + + The remaining elements @{command fix} and @{command assume} build up + a local context (see \secref{sec:framework-context}), while + @{command show} refines a pending sub-goal by the rule resulting + from a nested sub-proof (see \secref{sec:framework-subproof}). + Further derived concepts will support calculational reasoning (see + \secref{sec:framework-calc}). +*} + + +subsection {* Context elements \label{sec:framework-context} *} + +text {* + In judgments @{text "\ \ \"} of the primitive framework, @{text "\"} + essentially acts like a proof context. Isar elaborates this idea + towards a higher-level notion, with additional information for + type-inference, term abbreviations, local facts, hypotheses etc. + + The element @{command fix}~@{text "x :: \"} declares a local + parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in + results exported from the context, @{text "x"} may become anything. + The @{command assume}~@{text "\inference\"} element provides a + general interface to hypotheses: ``@{command assume}~@{text + "\inference\ A"}'' produces @{text "A \ A"} locally, while the + included inference tells how to discharge @{text A} from results + @{text "A \ B"} later on. There is no user-syntax for @{text + "\inference\"}, i.e.\ it may only occur internally when derived + commands are defined in ML. + + At the user-level, the default inference for @{command assume} is + @{inference discharge} as given below. The additional variants + @{command presume} and @{command def} are defined as follows: + + \medskip + \begin{tabular}{rcl} + @{command presume}~@{text A} & @{text "\"} & @{command assume}~@{text "\weak\discharge\ A"} \\ + @{command def}~@{text "x \ a"} & @{text "\"} & @{command fix}~@{text x}~@{command assume}~@{text "\expansion\ x \ a"} \\ + \end{tabular} + \medskip + + \[ + \infer[(@{inference_def discharge})]{@{text "\\ - A \ #A \ B"}}{@{text "\\ \ B"}} + \] + \[ + \infer[(@{inference_def "weak\discharge"})]{@{text "\\ - A \ A \ B"}}{@{text "\\ \ B"}} + \] + \[ + \infer[(@{inference_def expansion})]{@{text "\\ - (x \ a) \ B a"}}{@{text "\\ \ B x"}} + \] + + \medskip Note that @{inference discharge} and @{inference + "weak\discharge"} differ in the marker for @{prop A}, which is + relevant when the result of a @{command fix}-@{command + assume}-@{command show} outline is composed with a pending goal, + cf.\ \secref{sec:framework-subproof}. + + The most interesting derived context element in Isar is @{command + obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized + elimination steps in a purely forward manner. The @{command obtain} + command takes a specification of parameters @{text "\<^vec>x"} and + assumptions @{text "\<^vec>A"} to be added to the context, together + with a proof of a case rule stating that this extension is + conservative (i.e.\ may be removed from closed results later on): + + \medskip + \begin{tabular}{l} + @{text "\facts\"}~~@{command obtain}~@{text "\<^vec>x \ \<^vec>A \<^vec>x \proof\ \"} \\[0.5ex] + \quad @{command have}~@{text "case: \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis\"} \\ + \quad @{command proof}~@{method "-"} \\ + \qquad @{command fix}~@{text thesis} \\ + \qquad @{command assume}~@{text "[intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis"} \\ + \qquad @{command show}~@{text thesis}~@{command using}~@{text "\facts\ \proof\"} \\ + \quad @{command qed} \\ + \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\elimination case\ \<^vec>A \<^vec>x"} \\ + \end{tabular} + \medskip + + \[ + \infer[(@{inference elimination})]{@{text "\ \ B"}}{ + \begin{tabular}{rl} + @{text "case:"} & + @{text "\ \ \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis"} \\[0.2ex] + @{text "result:"} & + @{text "\ \ \<^vec>A \<^vec>y \ B"} \\[0.2ex] + \end{tabular}} + \] + + \noindent Here the name ``@{text thesis}'' is a specific convention + for an arbitrary-but-fixed proposition; in the primitive natural + deduction rules shown before we have occasionally used @{text C}. + The whole statement of ``@{command obtain}~@{text x}~@{keyword + "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"} + may be assumed for some arbitrary-but-fixed @{text "x"}. Also note + that ``@{command obtain}~@{text "A \ B"}'' without parameters + is similar to ``@{command have}~@{text "A \ B"}'', but the + latter involves multiple sub-goals. + + \medskip The subsequent Isar proof texts explain all context + elements introduced above using the formal proof language itself. + After finishing a local proof within a block, we indicate the + exported result via @{command note}. +*} + +(*<*) +theorem True +proof +(*>*) + txt_raw {* \begin{minipage}[t]{0.45\textwidth} *} + { + fix x + have "B x" sorry %noproof + } + note `\x. B x` + txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) + { + assume A + have B sorry %noproof + } + note `A \ B` + txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) + { + def x \ a + have "B x" sorry %noproof + } + note `B a` + txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) + { + obtain x where "A x" sorry %noproof + have B sorry %noproof + } + note `B` + txt_raw {* \end{minipage} *} +(*<*) +qed +(*>*) + +text {* + \bigskip\noindent This illustrates the meaning of Isar context + elements without goals getting in between. +*} + +subsection {* Structured statements \label{sec:framework-stmt} *} + +text {* + The category @{text "statement"} of top-level theorem specifications + is defined as follows: + + \medskip + \begin{tabular}{rcl} + @{text "statement"} & @{text "\"} & @{text "name: props \ \"} \\ + & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex] + + @{text "context"} & @{text "\"} & @{text "\ vars \ \"} \\ + & @{text "|"} & @{text "\ name: props \ \"} \\ + + @{text "conclusion"} & @{text "\"} & @{text "\ name: props \ \"} \\ + & @{text "|"} & @{text "\ vars \ \ \ name: props \ \"} \\ + & & \quad @{text "\ \"} \\ + \end{tabular} + + \medskip\noindent A simple @{text "statement"} consists of named + propositions. The full form admits local context elements followed + by the actual conclusions, such as ``@{keyword "fixes"}~@{text + x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B + x"}''. The final result emerges as a Pure rule after discharging + the context: @{prop "\x. A x \ B x"}. + + The @{keyword "obtains"} variant is another abbreviation defined + below; unlike @{command obtain} (cf.\ + \secref{sec:framework-context}) there may be several ``cases'' + separated by ``@{text "\"}'', each consisting of several + parameters (@{text "vars"}) and several premises (@{text "props"}). + This specifies multi-branch elimination rules. + + \medskip + \begin{tabular}{l} + @{text "\ \<^vec>x \ \<^vec>A \<^vec>x \ \ \"} \\[0.5ex] + \quad @{text "\ thesis"} \\ + \quad @{text "\ [intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis \ \"} \\ + \quad @{text "\ thesis"} \\ + \end{tabular} + \medskip + + Presenting structured statements in such an ``open'' format usually + simplifies the subsequent proof, because the outer structure of the + problem is already laid out directly. E.g.\ consider the following + canonical patterns for @{text "\"} and @{text "\"}, + respectively: +*} + +text_raw {*\begin{minipage}{0.5\textwidth}*} + +theorem + fixes x and y + assumes "A x" and "B y" + shows "C x y" +proof - + from `A x` and `B y` + show "C x y" sorry %noproof +qed + +text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} + +theorem + obtains x and y + where "A x" and "B y" +proof - + have "A a" and "B b" sorry %noproof + then show thesis .. +qed + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Here local facts \isacharbackquoteopen@{text "A + x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B + y"}\isacharbackquoteclose\ are referenced immediately; there is no + need to decompose the logical rule structure again. In the second + proof the final ``@{command then}~@{command show}~@{text + thesis}~@{command ".."}'' involves the local rule case @{text "\x + y. A x \ B y \ thesis"} for the particular instance of terms @{text + "a"} and @{text "b"} produced in the body. +*} + + +subsection {* Structured proof refinement \label{sec:framework-subproof} *} + +text {* + By breaking up the grammar for the Isar proof language, we may + understand a proof text as a linear sequence of individual proof + commands. These are interpreted as transitions of the Isar virtual + machine (Isar/VM), which operates on a block-structured + configuration in single steps. This allows users to write proof + texts in an incremental manner, and inspect intermediate + configurations for debugging. + + The basic idea is analogous to evaluating algebraic expressions on a + stack machine: @{text "(a + b) \ c"} then corresponds to a sequence + of single transitions for each symbol @{text "(, a, +, b, ), \, c"}. + In Isar the algebraic values are facts or goals, and the operations + are inferences. + + \medskip The Isar/VM state maintains a stack of nodes, each node + contains the local proof context, the linguistic mode, and a pending + goal (optional). The mode determines the type of transition that + may be performed next, it essentially alternates between forward and + backward reasoning, with an intermediate stage for chained facts + (see \figref{fig:isar-vm}). + + \begin{figure}[htb] + \begin{center} + \includegraphics[width=0.8\textwidth]{isar-vm} + \end{center} + \caption{Isar/VM modes}\label{fig:isar-vm} + \end{figure} + + For example, in @{text "state"} mode Isar acts like a mathematical + scratch-pad, accepting declarations like @{command fix}, @{command + assume}, and claims like @{command have}, @{command show}. A goal + statement changes the mode to @{text "prove"}, which means that we + may now refine the problem via @{command unfolding} or @{command + proof}. Then we are again in @{text "state"} mode of a proof body, + which may issue @{command show} statements to solve pending + sub-goals. A concluding @{command qed} will return to the original + @{text "state"} mode one level upwards. The subsequent Isar/VM + trace indicates block structure, linguistic mode, goal state, and + inferences: +*} + +text_raw {* \begingroup\footnotesize *} +(*<*)notepad begin +(*>*) + txt_raw {* \begin{minipage}[t]{0.18\textwidth} *} + have "A \ B" + proof + assume A + show B + sorry %noproof + qed + txt_raw {* \end{minipage}\quad +\begin{minipage}[t]{0.06\textwidth} +@{text "begin"} \\ +\\ +\\ +@{text "begin"} \\ +@{text "end"} \\ +@{text "end"} \\ +\end{minipage} +\begin{minipage}[t]{0.08\textwidth} +@{text "prove"} \\ +@{text "state"} \\ +@{text "state"} \\ +@{text "prove"} \\ +@{text "state"} \\ +@{text "state"} \\ +\end{minipage}\begin{minipage}[t]{0.35\textwidth} +@{text "(A \ B) \ #(A \ B)"} \\ +@{text "(A \ B) \ #(A \ B)"} \\ +\\ +\\ +@{text "#(A \ B)"} \\ +@{text "A \ B"} \\ +\end{minipage}\begin{minipage}[t]{0.4\textwidth} +@{text "(init)"} \\ +@{text "(resolution impI)"} \\ +\\ +\\ +@{text "(refinement #A \ B)"} \\ +@{text "(finish)"} \\ +\end{minipage} *} +(*<*) +end +(*>*) +text_raw {* \endgroup *} + +text {* + \noindent Here the @{inference refinement} inference from + \secref{sec:framework-resolution} mediates composition of Isar + sub-proofs nicely. Observe that this principle incorporates some + degree of freedom in proof composition. In particular, the proof + body allows parameters and assumptions to be re-ordered, or commuted + according to Hereditary Harrop Form. Moreover, context elements + that are not used in a sub-proof may be omitted altogether. For + example: +*} + +text_raw {*\begin{minipage}{0.5\textwidth}*} + +(*<*) +notepad +begin +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix x and y + assume "A x" and "B y" + show "C x y" sorry %noproof + qed + +txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} + +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix x assume "A x" + fix y assume "B y" + show "C x y" sorry %noproof + qed + +txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*} + +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix y assume "B y" + fix x assume "A x" + show "C x y" sorry + qed + +txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} +(*<*) +next +(*>*) + have "\x y. A x \ B y \ C x y" + proof - + fix y assume "B y" + fix x + show "C x y" sorry + qed +(*<*) +end +(*>*) + +text_raw {*\end{minipage}*} + +text {* + \medskip\noindent Such ``peephole optimizations'' of Isar texts are + practically important to improve readability, by rearranging + contexts elements according to the natural flow of reasoning in the + body, while still observing the overall scoping rules. + + \medskip This illustrates the basic idea of structured proof + processing in Isar. The main mechanisms are based on natural + deduction rule composition within the Pure framework. In + particular, there are no direct operations on goal states within the + proof body. Moreover, there is no hidden automated reasoning + involved, just plain unification. +*} + + +subsection {* Calculational reasoning \label{sec:framework-calc} *} + +text {* + The existing Isar infrastructure is sufficiently flexible to support + calculational reasoning (chains of transitivity steps) as derived + concept. The generic proof elements introduced below depend on + rules declared as @{attribute trans} in the context. It is left to + the object-logic to provide a suitable rule collection for mixed + relations of @{text "="}, @{text "<"}, @{text "\"}, @{text "\"}, + @{text "\"} etc. Due to the flexibility of rule composition + (\secref{sec:framework-resolution}), substitution of equals by + equals is covered as well, even substitution of inequalities + involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD} + and \cite{Bauer-Wenzel:2001}. + + The generic calculational mechanism is based on the observation that + rules such as @{text "trans:"}~@{prop "x = y \ y = z \ x = z"} + proceed from the premises towards the conclusion in a deterministic + fashion. Thus we may reason in forward mode, feeding intermediate + results into rules selected from the context. The course of + reasoning is organized by maintaining a secondary fact called + ``@{fact calculation}'', apart from the primary ``@{fact this}'' + already provided by the Isar primitives. In the definitions below, + @{attribute OF} refers to @{inference resolution} + (\secref{sec:framework-resolution}) with multiple rule arguments, + and @{text "trans"} represents to a suitable rule from the context: + + \begin{matharray}{rcl} + @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\ + @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex] + @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\ + \end{matharray} + + \noindent The start of a calculation is determined implicitly in the + text: here @{command also} sets @{fact calculation} to the current + result; any subsequent occurrence will update @{fact calculation} by + combination with the next result and a transitivity rule. The + calculational sequence is concluded via @{command finally}, where + the final result is exposed for use in a concluding claim. + + Here is a canonical proof pattern, using @{command have} to + establish the intermediate results: +*} + +(*<*) +notepad +begin +(*>*) + have "a = b" sorry + also have "\ = c" sorry + also have "\ = d" sorry + finally have "a = d" . +(*<*) +end +(*>*) + +text {* + \noindent The term ``@{text "\"}'' above is a special abbreviation + provided by the Isabelle/Isar syntax layer: it statically refers to + the right-hand side argument of the previous statement given in the + text. Thus it happens to coincide with relevant sub-expressions in + the calculational chain, but the exact correspondence is dependent + on the transitivity rules being involved. + + \medskip Symmetry rules such as @{prop "x = y \ y = x"} are like + transitivities with only one premise. Isar maintains a separate + rule collection declared via the @{attribute sym} attribute, to be + used in fact expressions ``@{text "a [symmetric]"}'', or single-step + proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command + have}~@{text "y = x"}~@{command ".."}''. +*} + +end \ No newline at end of file diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Generic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Generic.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1279 @@ +theory Generic +imports Base Main +begin + +chapter {* Generic tools and packages \label{ch:gen-tools} *} + +section {* Configuration options \label{sec:config} *} + +text {* Isabelle/Pure maintains a record of named configuration + options within the theory or proof context, with values of type + @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type + string}. Tools may declare options in ML, and then refer to these + values (relative to the context). Thus global reference variables + are easily avoided. The user may change the value of a + configuration option by means of an associated attribute of the same + name. This form of context declaration works particularly well with + commands such as @{command "declare"} or @{command "using"} like + this: +*} + +declare [[show_main_goal = false]] + +notepad +begin + note [[show_main_goal = true]] +end + +text {* For historical reasons, some tools cannot take the full proof + context into account and merely refer to the background theory. + This is accommodated by configuration options being declared as + ``global'', which may not be changed within a local context. + + \begin{matharray}{rcll} + @{command_def "print_configs"} & : & @{text "context \"} \\ + \end{matharray} + + @{rail " + @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? + "} + + \begin{description} + + \item @{command "print_configs"} prints the available configuration + options, with names, types, and current values. + + \item @{text "name = value"} as an attribute expression modifies the + named option, with the syntax of the value depending on the option's + type. For @{ML_type bool} the default value is @{text true}. Any + attempt to change a global option in a local context is ignored. + + \end{description} +*} + + +section {* Basic proof tools *} + +subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *} + +text {* + \begin{matharray}{rcl} + @{method_def unfold} & : & @{text method} \\ + @{method_def fold} & : & @{text method} \\ + @{method_def insert} & : & @{text method} \\[0.5ex] + @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def intro} & : & @{text method} \\ + @{method_def elim} & : & @{text method} \\ + @{method_def succeed} & : & @{text method} \\ + @{method_def fail} & : & @{text method} \\ + \end{matharray} + + @{rail " + (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs} + ; + (@@{method erule} | @@{method drule} | @@{method frule}) + ('(' @{syntax nat} ')')? @{syntax thmrefs} + ; + (@@{method intro} | @@{method elim}) @{syntax thmrefs}? + "} + + \begin{description} + + \item @{method unfold}~@{text "a\<^sub>1 \ a\<^sub>n"} and @{method fold}~@{text + "a\<^sub>1 \ a\<^sub>n"} expand (or fold back) the given definitions throughout + all goals; any chained facts provided are inserted into the goal and + subject to rewriting as well. + + \item @{method insert}~@{text "a\<^sub>1 \ a\<^sub>n"} inserts theorems as facts + into all goals of the proof state. Note that current facts + indicated for forward chaining are ignored. + + \item @{method erule}~@{text "a\<^sub>1 \ a\<^sub>n"}, @{method + drule}~@{text "a\<^sub>1 \ a\<^sub>n"}, and @{method frule}~@{text + "a\<^sub>1 \ a\<^sub>n"} are similar to the basic @{method rule} + method (see \secref{sec:pure-meth-att}), but apply rules by + elim-resolution, destruct-resolution, and forward-resolution, + respectively \cite{isabelle-implementation}. The optional natural + number argument (default 0) specifies additional assumption steps to + be performed here. + + Note that these methods are improper ones, mainly serving for + experimentation and tactic script emulation. Different modes of + basic rule application are usually expressed in Isar at the proof + language level, rather than via implicit proof state manipulations. + For example, a proper single-step elimination would be done using + the plain @{method rule} method, with forward chaining of current + facts. + + \item @{method intro} and @{method elim} repeatedly refine some goal + by intro- or elim-resolution, after having inserted any chained + facts. Exactly the rules given as arguments are taken into account; + this allows fine-tuned decomposition of a proof problem, in contrast + to common automated tools. + + \item @{method succeed} yields a single (unchanged) result; it is + the identity of the ``@{text ","}'' method combinator (cf.\ + \secref{sec:proof-meth}). + + \item @{method fail} yields an empty result sequence; it is the + identity of the ``@{text "|"}'' method combinator (cf.\ + \secref{sec:proof-meth}). + + \end{description} + + \begin{matharray}{rcl} + @{attribute_def tagged} & : & @{text attribute} \\ + @{attribute_def untagged} & : & @{text attribute} \\[0.5ex] + @{attribute_def THEN} & : & @{text attribute} \\ + @{attribute_def unfolded} & : & @{text attribute} \\ + @{attribute_def folded} & : & @{text attribute} \\ + @{attribute_def abs_def} & : & @{text attribute} \\[0.5ex] + @{attribute_def rotated} & : & @{text attribute} \\ + @{attribute_def (Pure) elim_format} & : & @{text attribute} \\ + @{attribute_def standard}@{text "\<^sup>*"} & : & @{text attribute} \\ + @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{attribute tagged} @{syntax name} @{syntax name} + ; + @@{attribute untagged} @{syntax name} + ; + @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thmref} + ; + (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs} + ; + @@{attribute rotated} @{syntax int}? + "} + + \begin{description} + + \item @{attribute tagged}~@{text "name value"} and @{attribute + untagged}~@{text name} add and remove \emph{tags} of some theorem. + Tags may be any list of string pairs that serve as formal comment. + The first string is considered the tag name, the second its value. + Note that @{attribute untagged} removes any tags of the same name. + + \item @{attribute THEN}~@{text a} composes rules by resolution; it + resolves with the first premise of @{text a} (an alternative + position may be also specified). See also @{ML_op "RS"} in + \cite{isabelle-implementation}. + + \item @{attribute unfolded}~@{text "a\<^sub>1 \ a\<^sub>n"} and @{attribute + folded}~@{text "a\<^sub>1 \ a\<^sub>n"} expand and fold back again the given + definitions throughout a rule. + + \item @{attribute abs_def} turns an equation of the form @{prop "f x + y \ t"} into @{prop "f \ \x y. t"}, which ensures that @{method + simp} or @{method unfold} steps always expand it. This also works + for object-logic equality. + + \item @{attribute rotated}~@{text n} rotate the premises of a + theorem by @{text n} (default 1). + + \item @{attribute (Pure) elim_format} turns a destruction rule into + elimination rule format, by resolving with the rule @{prop "PROP A \ + (PROP A \ PROP B) \ PROP B"}. + + Note that the Classical Reasoner (\secref{sec:classical}) provides + its own version of this operation. + + \item @{attribute standard} puts a theorem into the standard form of + object-rules at the outermost theory level. Note that this + operation violates the local proof context (including active + locales). + + \item @{attribute no_vars} replaces schematic variables by free + ones; this is mainly for tuning output of pretty printed theorems. + + \end{description} +*} + + +subsection {* Low-level equational reasoning *} + +text {* + \begin{matharray}{rcl} + @{method_def subst} & : & @{text method} \\ + @{method_def hypsubst} & : & @{text method} \\ + @{method_def split} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method subst} ('(' 'asm' ')')? \\ ('(' (@{syntax nat}+) ')')? @{syntax thmref} + ; + @@{method split} @{syntax thmrefs} + "} + + These methods provide low-level facilities for equational reasoning + that are intended for specialized applications only. Normally, + single step calculations would be performed in a structured text + (see also \secref{sec:calculation}), while the Simplifier methods + provide the canonical way for automated normalization (see + \secref{sec:simplifier}). + + \begin{description} + + \item @{method subst}~@{text eq} performs a single substitution step + using rule @{text eq}, which may be either a meta or object + equality. + + \item @{method subst}~@{text "(asm) eq"} substitutes in an + assumption. + + \item @{method subst}~@{text "(i \ j) eq"} performs several + substitutions in the conclusion. The numbers @{text i} to @{text j} + indicate the positions to substitute at. Positions are ordered from + the top of the term tree moving down from left to right. For + example, in @{text "(a + b) + (c + d)"} there are three positions + where commutativity of @{text "+"} is applicable: 1 refers to @{text + "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}. + + If the positions in the list @{text "(i \ j)"} are non-overlapping + (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may + assume all substitutions are performed simultaneously. Otherwise + the behaviour of @{text subst} is not specified. + + \item @{method subst}~@{text "(asm) (i \ j) eq"} performs the + substitutions in the assumptions. The positions refer to the + assumptions in order from left to right. For example, given in a + goal of the form @{text "P (a + b) \ P (c + d) \ \"}, position 1 of + commutativity of @{text "+"} is the subterm @{text "a + b"} and + position 2 is the subterm @{text "c + d"}. + + \item @{method hypsubst} performs substitution using some + assumption; this only works for equations of the form @{text "x = + t"} where @{text x} is a free or bound variable. + + \item @{method split}~@{text "a\<^sub>1 \ a\<^sub>n"} performs single-step case + splitting using the given rules. Splitting is performed in the + conclusion or some assumption of the subgoal, depending of the + structure of the rule. + + Note that the @{method simp} method already involves repeated + application of split rules as declared in the current context, using + @{attribute split}, for example. + + \end{description} +*} + + +subsection {* Further tactic emulations \label{sec:tactics} *} + +text {* + The following improper proof methods emulate traditional tactics. + These admit direct access to the goal state, which is normally + considered harmful! In particular, this may involve both numbered + goal addressing (default 1), and dynamic instantiation within the + scope of some subgoal. + + \begin{warn} + Dynamic instantiations refer to universally quantified parameters + of a subgoal (the dynamic context) rather than fixed variables and + term abbreviations of a (static) Isar context. + \end{warn} + + Tactic emulation methods, unlike their ML counterparts, admit + simultaneous instantiation from both dynamic and static contexts. + If names occur in both contexts goal parameters hide locally fixed + variables. Likewise, schematic variables refer to term + abbreviations, if present in the static context. Otherwise the + schematic variable is interpreted as a schematic variable and left + to be solved by unification with certain parts of the subgoal. + + Note that the tactic emulation proof methods in Isabelle/Isar are + consistently named @{text foo_tac}. Note also that variable names + occurring on left hand sides of instantiations must be preceded by a + question mark if they coincide with a keyword or contain dots. This + is consistent with the attribute @{attribute "where"} (see + \secref{sec:pure-meth-att}). + + \begin{matharray}{rcl} + @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\ + \end{matharray} + + @{rail " + (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} | + @@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \\ + ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} ) + ; + @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +) + ; + @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +) + ; + @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}? + ; + (@@{method tactic} | @@{method raw_tactic}) @{syntax text} + ; + + dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and') + "} + +\begin{description} + + \item @{method rule_tac} etc. do resolution of rules with explicit + instantiation. This works the same way as the ML tactics @{ML + res_inst_tac} etc. (see \cite{isabelle-implementation}) + + Multiple rules may be only given if there is no instantiation; then + @{method rule_tac} is the same as @{ML resolve_tac} in ML (see + \cite{isabelle-implementation}). + + \item @{method cut_tac} inserts facts into the proof state as + assumption of a subgoal; instantiations may be given as well. Note + that the scope of schematic variables is spread over the main goal + statement and rule premises are turned into new subgoals. This is + in contrast to the regular method @{method insert} which inserts + closed rule statements. + + \item @{method thin_tac}~@{text \} deletes the specified premise + from a subgoal. Note that @{text \} may contain schematic + variables, to abbreviate the intended proposition; the first + matching subgoal premise will be deleted. Removing useless premises + from a subgoal increases its readability and can make search tactics + run faster. + + \item @{method subgoal_tac}~@{text "\\<^sub>1 \ \\<^sub>n"} adds the propositions + @{text "\\<^sub>1 \ \\<^sub>n"} as local premises to a subgoal, and poses the same + as new subgoals (in the original context). + + \item @{method rename_tac}~@{text "x\<^sub>1 \ x\<^sub>n"} renames parameters of a + goal according to the list @{text "x\<^sub>1, \, x\<^sub>n"}, which refers to the + \emph{suffix} of variables. + + \item @{method rotate_tac}~@{text n} rotates the premises of a + subgoal by @{text n} positions: from right to left if @{text n} is + positive, and from left to right if @{text n} is negative; the + default value is 1. + + \item @{method tactic}~@{text "text"} produces a proof method from + any ML text of type @{ML_type tactic}. Apart from the usual ML + environment and the current proof context, the ML code may refer to + the locally bound values @{ML_text facts}, which indicates any + current facts used for forward-chaining. + + \item @{method raw_tactic} is similar to @{method tactic}, but + presents the goal state in its raw internal form, where simultaneous + subgoals appear as conjunction of the logical framework instead of + the usual split into several subgoals. While feature this is useful + for debugging of complex method definitions, it should not never + appear in production theories. + + \end{description} +*} + + +section {* The Simplifier \label{sec:simplifier} *} + +subsection {* Simplification methods *} + +text {* + \begin{matharray}{rcl} + @{method_def simp} & : & @{text method} \\ + @{method_def simp_all} & : & @{text method} \\ + \end{matharray} + + @{rail " + (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) + ; + + opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' + ; + @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') | + 'split' (() | 'add' | 'del')) ':' @{syntax thmrefs} + "} + + \begin{description} + + \item @{method simp} invokes the Simplifier, after declaring + additional rules according to the arguments given. Note that the + @{text only} modifier first removes all other rewrite rules, + congruences, and looper tactics (including splits), and then behaves + like @{text add}. + + \medskip The @{text cong} modifiers add or delete Simplifier + congruence rules (see also \secref{sec:simp-cong}), the default is + to add. + + \medskip The @{text split} modifiers add or delete rules for the + Splitter (see also \cite{isabelle-ref}), the default is to add. + This works only if the Simplifier method has been properly setup to + include the Splitter (all major object logics such HOL, HOLCF, FOL, + ZF do this already). + + \item @{method simp_all} is similar to @{method simp}, but acts on + all goals (backwards from the last to the first one). + + \end{description} + + By default the Simplifier methods take local assumptions fully into + account, using equational assumptions in the subsequent + normalization process, or simplifying assumptions themselves (cf.\ + @{ML asm_full_simp_tac} in \cite{isabelle-ref}). In structured + proofs this is usually quite well behaved in practice: just the + local premises of the actual goal are involved, additional facts may + be inserted via explicit forward-chaining (via @{command "then"}, + @{command "from"}, @{command "using"} etc.). + + Additional Simplifier options may be specified to tune the behavior + further (mostly for unstructured scripts with many accidental local + facts): ``@{text "(no_asm)"}'' means assumptions are ignored + completely (cf.\ @{ML simp_tac}), ``@{text "(no_asm_simp)"}'' means + assumptions are used in the simplification of the conclusion but are + not themselves simplified (cf.\ @{ML asm_simp_tac}), and ``@{text + "(no_asm_use)"}'' means assumptions are simplified but are not used + in the simplification of each other or the conclusion (cf.\ @{ML + full_simp_tac}). For compatibility reasons, there is also an option + ``@{text "(asm_lr)"}'', which means that an assumption is only used + for simplifying assumptions which are to the right of it (cf.\ @{ML + asm_lr_simp_tac}). + + The configuration option @{text "depth_limit"} limits the number of + recursive invocations of the simplifier during conditional + rewriting. + + \medskip The Splitter package is usually configured to work as part + of the Simplifier. The effect of repeatedly applying @{ML + split_tac} can be simulated by ``@{text "(simp only: split: + a\<^sub>1 \ a\<^sub>n)"}''. There is also a separate @{text split} + method available for single-step case splitting. +*} + + +subsection {* Declaring rules *} + +text {* + \begin{matharray}{rcl} + @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def simp} & : & @{text attribute} \\ + @{attribute_def split} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + (@@{attribute simp} | @@{attribute split}) (() | 'add' | 'del') + "} + + \begin{description} + + \item @{command "print_simpset"} prints the collection of rules + declared to the Simplifier, which is also known as ``simpset'' + internally \cite{isabelle-ref}. + + \item @{attribute simp} declares simplification rules. + + \item @{attribute split} declares case split rules. + + \end{description} +*} + + +subsection {* Congruence rules\label{sec:simp-cong} *} + +text {* + \begin{matharray}{rcl} + @{attribute_def cong} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{attribute cong} (() | 'add' | 'del') + "} + + \begin{description} + + \item @{attribute cong} declares congruence rules to the Simplifier + context. + + \end{description} + + Congruence rules are equalities of the form @{text [display] + "\ \ f ?x\<^sub>1 \ ?x\<^sub>n = f ?y\<^sub>1 \ ?y\<^sub>n"} + + This controls the simplification of the arguments of @{text f}. For + example, some arguments can be simplified under additional + assumptions: @{text [display] "?P\<^sub>1 \ ?Q\<^sub>1 \ (?Q\<^sub>1 \ ?P\<^sub>2 \ ?Q\<^sub>2) \ + (?P\<^sub>1 \ ?P\<^sub>2) \ (?Q\<^sub>1 \ ?Q\<^sub>2)"} + + Given this rule, the simplifier assumes @{text "?Q\<^sub>1"} and extracts + rewrite rules from it when simplifying @{text "?P\<^sub>2"}. Such local + assumptions are effective for rewriting formulae such as @{text "x = + 0 \ y + x = y"}. + + %FIXME + %The local assumptions are also provided as theorems to the solver; + %see \secref{sec:simp-solver} below. + + \medskip The following congruence rule for bounded quantifiers also + supplies contextual information --- about the bound variable: + @{text [display] "(?A = ?B) \ (\x. x \ ?B \ ?P x \ ?Q x) \ + (\x \ ?A. ?P x) \ (\x \ ?B. ?Q x)"} + + \medskip This congruence rule for conditional expressions can + supply contextual information for simplifying the arms: + @{text [display] "?p = ?q \ (?q \ ?a = ?c) \ (\ ?q \ ?b = ?d) \ + (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} + + A congruence rule can also \emph{prevent} simplification of some + arguments. Here is an alternative congruence rule for conditional + expressions that conforms to non-strict functional evaluation: + @{text [display] "?p = ?q \ (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} + + Only the first argument is simplified; the others remain unchanged. + This can make simplification much faster, but may require an extra + case split over the condition @{text "?q"} to prove the goal. +*} + + +subsection {* Simplification procedures *} + +text {* Simplification procedures are ML functions that produce proven + rewrite rules on demand. They are associated with higher-order + patterns that approximate the left-hand sides of equations. The + Simplifier first matches the current redex against one of the LHS + patterns; if this succeeds, the corresponding ML function is + invoked, passing the Simplifier context and redex term. Thus rules + may be specifically fashioned for particular situations, resulting + in a more powerful mechanism than term rewriting by a fixed set of + rules. + + Any successful result needs to be a (possibly conditional) rewrite + rule @{text "t \ u"} that is applicable to the current redex. The + rule will be applied just as any ordinary rewrite rule. It is + expected to be already in \emph{internal form}, bypassing the + automatic preprocessing of object-level equivalences. + + \begin{matharray}{rcl} + @{command_def "simproc_setup"} & : & @{text "local_theory \ local_theory"} \\ + simproc & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' + @{syntax text} \\ (@'identifier' (@{syntax nameref}+))? + ; + + @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) + "} + + \begin{description} + + \item @{command "simproc_setup"} defines a named simplification + procedure that is invoked by the Simplifier whenever any of the + given term patterns match the current redex. The implementation, + which is provided as ML source text, needs to be of type @{ML_type + "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type + cterm} represents the current redex @{text r} and the result is + supposed to be some proven rewrite rule @{text "r \ r'"} (or a + generalized version), or @{ML NONE} to indicate failure. The + @{ML_type simpset} argument holds the full context of the current + Simplifier invocation, including the actual Isar proof context. The + @{ML_type morphism} informs about the difference of the original + compilation context wrt.\ the one of the actual application later + on. The optional @{keyword "identifier"} specifies theorems that + represent the logical content of the abstract theory of this + simproc. + + Morphisms and identifiers are only relevant for simprocs that are + defined within a local target context, e.g.\ in a locale. + + \item @{text "simproc add: name"} and @{text "simproc del: name"} + add or delete named simprocs to the current Simplifier context. The + default is to add a simproc. Note that @{command "simproc_setup"} + already adds the new simproc to the subsequent context. + + \end{description} +*} + + +subsubsection {* Example *} + +text {* The following simplification procedure for @{thm + [source=false, show_types] unit_eq} in HOL performs fine-grained + control over rule application, beyond higher-order pattern matching. + Declaring @{thm unit_eq} as @{attribute simp} directly would make + the simplifier loop! Note that a version of this simplification + procedure is already active in Isabelle/HOL. *} + +simproc_setup unit ("x::unit") = {* + fn _ => fn _ => fn ct => + if HOLogic.is_unit (term_of ct) then NONE + else SOME (mk_meta_eq @{thm unit_eq}) +*} + +text {* Since the Simplifier applies simplification procedures + frequently, it is important to make the failure check in ML + reasonably fast. *} + + +subsection {* Forward simplification *} + +text {* + \begin{matharray}{rcl} + @{attribute_def simplified} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{attribute simplified} opt? @{syntax thmrefs}? + ; + + opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' + "} + + \begin{description} + + \item @{attribute simplified}~@{text "a\<^sub>1 \ a\<^sub>n"} causes a theorem to + be simplified, either by exactly the specified rules @{text "a\<^sub>1, \, + a\<^sub>n"}, or the implicit Simplifier context if no arguments are given. + The result is fully simplified by default, including assumptions and + conclusion; the options @{text no_asm} etc.\ tune the Simplifier in + the same way as the for the @{text simp} method. + + Note that forward simplification restricts the simplifier to its + most basic operation of term rewriting; solver and looper tactics + \cite{isabelle-ref} are \emph{not} involved here. The @{text + simplified} attribute should be only rarely required under normal + circumstances. + + \end{description} +*} + + +section {* The Classical Reasoner \label{sec:classical} *} + +subsection {* Basic concepts *} + +text {* Although Isabelle is generic, many users will be working in + some extension of classical first-order logic. Isabelle/ZF is built + upon theory FOL, while Isabelle/HOL conceptually contains + first-order logic as a fragment. Theorem-proving in predicate logic + is undecidable, but many automated strategies have been developed to + assist in this task. + + Isabelle's classical reasoner is a generic package that accepts + certain information about a logic and delivers a suite of automatic + proof tools, based on rules that are classified and declared in the + context. These proof procedures are slow and simplistic compared + with high-end automated theorem provers, but they can save + considerable time and effort in practice. They can prove theorems + such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few + milliseconds (including full proof reconstruction): *} + +lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" + by blast + +lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" + by blast + +text {* The proof tools are generic. They are not restricted to + first-order logic, and have been heavily used in the development of + the Isabelle/HOL library and applications. The tactics can be + traced, and their components can be called directly; in this manner, + any proof can be viewed interactively. *} + + +subsubsection {* The sequent calculus *} + +text {* Isabelle supports natural deduction, which is easy to use for + interactive proof. But natural deduction does not easily lend + itself to automation, and has a bias towards intuitionism. For + certain proofs in classical logic, it can not be called natural. + The \emph{sequent calculus}, a generalization of natural deduction, + is easier to automate. + + A \textbf{sequent} has the form @{text "\ \ \"}, where @{text "\"} + and @{text "\"} are sets of formulae.\footnote{For first-order + logic, sequents can equivalently be made from lists or multisets of + formulae.} The sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} is + \textbf{valid} if @{text "P\<^sub>1 \ \ \ P\<^sub>m"} implies @{text "Q\<^sub>1 \ \ \ + Q\<^sub>n"}. Thus @{text "P\<^sub>1, \, P\<^sub>m"} represent assumptions, each of which + is true, while @{text "Q\<^sub>1, \, Q\<^sub>n"} represent alternative goals. A + sequent is \textbf{basic} if its left and right sides have a common + formula, as in @{text "P, Q \ Q, R"}; basic sequents are trivially + valid. + + Sequent rules are classified as \textbf{right} or \textbf{left}, + indicating which side of the @{text "\"} symbol they operate on. + Rules that operate on the right side are analogous to natural + deduction's introduction rules, and left rules are analogous to + elimination rules. The sequent calculus analogue of @{text "(\I)"} + is the rule + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "P, \ \ \, Q"}} + \] + Applying the rule backwards, this breaks down some implication on + the right side of a sequent; @{text "\"} and @{text "\"} stand for + the sets of formulae that are unaffected by the inference. The + analogue of the pair @{text "(\I1)"} and @{text "(\I2)"} is the + single rule + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "\ \ \, P, Q"}} + \] + This breaks down some disjunction on the right side, replacing it by + both disjuncts. Thus, the sequent calculus is a kind of + multiple-conclusion logic. + + To illustrate the use of multiple formulae on the right, let us + prove the classical theorem @{text "(P \ Q) \ (Q \ P)"}. Working + backwards, we reduce this formula to a basic sequent: + \[ + \infer[@{text "(\R)"}]{@{text "\ (P \ Q) \ (Q \ P)"}} + {\infer[@{text "(\R)"}]{@{text "\ (P \ Q), (Q \ P)"}} + {\infer[@{text "(\R)"}]{@{text "P \ Q, (Q \ P)"}} + {@{text "P, Q \ Q, P"}}}} + \] + + This example is typical of the sequent calculus: start with the + desired theorem and apply rules backwards in a fairly arbitrary + manner. This yields a surprisingly effective proof procedure. + Quantifiers add only few complications, since Isabelle handles + parameters and schematic variables. See \cite[Chapter + 10]{paulson-ml2} for further discussion. *} + + +subsubsection {* Simulating sequents by natural deduction *} + +text {* Isabelle can represent sequents directly, as in the + object-logic LK. But natural deduction is easier to work with, and + most object-logics employ it. Fortunately, we can simulate the + sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} by the Isabelle formula + @{text "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1"} where the order of + the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary. + Elim-resolution plays a key role in simulating sequent proofs. + + We can easily handle reasoning on the left. Elim-resolution with + the rules @{text "(\E)"}, @{text "(\E)"} and @{text "(\E)"} achieves + a similar effect as the corresponding sequent rules. For the other + connectives, we use sequent-style elimination rules instead of + destruction rules such as @{text "(\E1, 2)"} and @{text "(\E)"}. + But note that the rule @{text "(\L)"} has no effect under our + representation of sequents! + \[ + \infer[@{text "(\L)"}]{@{text "\ P, \ \ \"}}{@{text "\ \ \, P"}} + \] + + What about reasoning on the right? Introduction rules can only + affect the formula in the conclusion, namely @{text "Q\<^sub>1"}. The + other right-side formulae are represented as negated assumptions, + @{text "\ Q\<^sub>2, \, \ Q\<^sub>n"}. In order to operate on one of these, it + must first be exchanged with @{text "Q\<^sub>1"}. Elim-resolution with the + @{text swap} rule has this effect: @{text "\ P \ (\ R \ P) \ R"} + + To ensure that swaps occur only when necessary, each introduction + rule is converted into a swapped form: it is resolved with the + second premise of @{text "(swap)"}. The swapped form of @{text + "(\I)"}, which might be called @{text "(\\E)"}, is + @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} + + Similarly, the swapped form of @{text "(\I)"} is + @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} + + Swapped introduction rules are applied using elim-resolution, which + deletes the negated formula. Our representation of sequents also + requires the use of ordinary introduction rules. If we had no + regard for readability of intermediate goal states, we could treat + the right side more uniformly by representing sequents as @{text + [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} +*} + + +subsubsection {* Extra rules for the sequent calculus *} + +text {* As mentioned, destruction rules such as @{text "(\E1, 2)"} and + @{text "(\E)"} must be replaced by sequent-style elimination rules. + In addition, we need rules to embody the classical equivalence + between @{text "P \ Q"} and @{text "\ P \ Q"}. The introduction + rules @{text "(\I1, 2)"} are replaced by a rule that simulates + @{text "(\R)"}: @{text [display] "(\ Q \ P) \ P \ Q"} + + The destruction rule @{text "(\E)"} is replaced by @{text [display] + "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} + + Quantifier replication also requires special rules. In classical + logic, @{text "\x. P x"} is equivalent to @{text "\ (\x. \ P x)"}; + the rules @{text "(\R)"} and @{text "(\L)"} are dual: + \[ + \infer[@{text "(\R)"}]{@{text "\ \ \, \x. P x"}}{@{text "\ \ \, \x. P x, P t"}} + \qquad + \infer[@{text "(\L)"}]{@{text "\x. P x, \ \ \"}}{@{text "P t, \x. P x, \ \ \"}} + \] + Thus both kinds of quantifier may be replicated. Theorems requiring + multiple uses of a universal formula are easy to invent; consider + @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any + @{text "n > 1"}. Natural examples of the multiple use of an + existential formula are rare; a standard one is @{text "\x. \y. P x + \ P y"}. + + Forgoing quantifier replication loses completeness, but gains + decidability, since the search space becomes finite. Many useful + theorems can be proved without replication, and the search generally + delivers its verdict in a reasonable time. To adopt this approach, + represent the sequent rules @{text "(\R)"}, @{text "(\L)"} and + @{text "(\R)"} by @{text "(\I)"}, @{text "(\E)"} and @{text "(\I)"}, + respectively, and put @{text "(\E)"} into elimination form: @{text + [display] "\x. P x \ (P t \ Q) \ Q"} + + Elim-resolution with this rule will delete the universal formula + after a single use. To replicate universal quantifiers, replace the + rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} + + To replicate existential quantifiers, replace @{text "(\I)"} by + @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} + + All introduction rules mentioned above are also useful in swapped + form. + + Replication makes the search space infinite; we must apply the rules + with care. The classical reasoner distinguishes between safe and + unsafe rules, applying the latter only when there is no alternative. + Depth-first search may well go down a blind alley; best-first search + is better behaved in an infinite search space. However, quantifier + replication is too expensive to prove any but the simplest theorems. +*} + + +subsection {* Rule declarations *} + +text {* The proof tools of the Classical Reasoner depend on + collections of rules declared in the context, which are classified + as introduction, elimination or destruction and as \emph{safe} or + \emph{unsafe}. In general, safe rules can be attempted blindly, + while unsafe rules must be used with care. A safe rule must never + reduce a provable goal to an unprovable set of subgoals. + + The rule @{text "P \ P \ Q"} is unsafe because it reduces @{text "P + \ Q"} to @{text "P"}, which might turn out as premature choice of an + unprovable subgoal. Any rule is unsafe whose premises contain new + unknowns. The elimination rule @{text "\x. P x \ (P t \ Q) \ Q"} is + unsafe, since it is applied via elim-resolution, which discards the + assumption @{text "\x. P x"} and replaces it by the weaker + assumption @{text "P t"}. The rule @{text "P t \ \x. P x"} is + unsafe for similar reasons. The quantifier duplication rule @{text + "\x. P x \ (P t \ \x. P x \ Q) \ Q"} is unsafe in a different sense: + since it keeps the assumption @{text "\x. P x"}, it is prone to + looping. In classical first-order logic, all rules are safe except + those mentioned above. + + The safe~/ unsafe distinction is vague, and may be regarded merely + as a way of giving some rules priority over others. One could argue + that @{text "(\E)"} is unsafe, because repeated application of it + could generate exponentially many subgoals. Induction rules are + unsafe because inductive proofs are difficult to set up + automatically. Any inference is unsafe that instantiates an unknown + in the proof state --- thus matching must be used, rather than + unification. Even proof by assumption is unsafe if it instantiates + unknowns shared with other subgoals. + + \begin{matharray}{rcl} + @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def intro} & : & @{text attribute} \\ + @{attribute_def elim} & : & @{text attribute} \\ + @{attribute_def dest} & : & @{text attribute} \\ + @{attribute_def rule} & : & @{text attribute} \\ + @{attribute_def iff} & : & @{text attribute} \\ + @{attribute_def swapped} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? + ; + @@{attribute rule} 'del' + ; + @@{attribute iff} (((() | 'add') '?'?) | 'del') + "} + + \begin{description} + + \item @{command "print_claset"} prints the collection of rules + declared to the Classical Reasoner, i.e.\ the @{ML_type claset} + within the context. + + \item @{attribute intro}, @{attribute elim}, and @{attribute dest} + declare introduction, elimination, and destruction rules, + respectively. By default, rules are considered as \emph{unsafe} + (i.e.\ not applied blindly without backtracking), while ``@{text + "!"}'' classifies as \emph{safe}. Rule declarations marked by + ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\ + \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps + of the @{method rule} method). The optional natural number + specifies an explicit weight argument, which is ignored by the + automated reasoning tools, but determines the search order of single + rule steps. + + Introduction rules are those that can be applied using ordinary + resolution. Their swapped forms are generated internally, which + will be applied using elim-resolution. Elimination rules are + applied using elim-resolution. Rules are sorted by the number of + new subgoals they will yield; rules that generate the fewest + subgoals will be tried first. Otherwise, later declarations take + precedence over earlier ones. + + Rules already present in the context with the same classification + are ignored. A warning is printed if the rule has already been + added with some other classification, but the rule is added anyway + as requested. + + \item @{attribute rule}~@{text del} deletes all occurrences of a + rule from the classical context, regardless of its classification as + introduction~/ elimination~/ destruction and safe~/ unsafe. + + \item @{attribute iff} declares logical equivalences to the + Simplifier and the Classical reasoner at the same time. + Non-conditional rules result in a safe introduction and elimination + pair; conditional ones are considered unsafe. Rules with negative + conclusion are automatically inverted (using @{text "\"}-elimination + internally). + + The ``@{text "?"}'' version of @{attribute iff} declares rules to + the Isabelle/Pure context only, and omits the Simplifier + declaration. + + \item @{attribute swapped} turns an introduction rule into an + elimination, by resolving with the classical swap principle @{text + "\ P \ (\ R \ P) \ R"} in the second position. This is mainly for + illustrative purposes: the Classical Reasoner already swaps rules + internally as explained above. + + \end{description} +*} + + +subsection {* Structured methods *} + +text {* + \begin{matharray}{rcl} + @{method_def rule} & : & @{text method} \\ + @{method_def contradiction} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method rule} @{syntax thmrefs}? + "} + + \begin{description} + + \item @{method rule} as offered by the Classical Reasoner is a + refinement over the Pure one (see \secref{sec:pure-meth-att}). Both + versions work the same, but the classical version observes the + classical rule context in addition to that of Isabelle/Pure. + + Common object logics (HOL, ZF, etc.) declare a rich collection of + classical rules (even if these would qualify as intuitionistic + ones), but only few declarations to the rule context of + Isabelle/Pure (\secref{sec:pure-meth-att}). + + \item @{method contradiction} solves some goal by contradiction, + deriving any result from both @{text "\ A"} and @{text A}. Chained + facts, which are guaranteed to participate, may appear in either + order. + + \end{description} +*} + + +subsection {* Automated methods *} + +text {* + \begin{matharray}{rcl} + @{method_def blast} & : & @{text method} \\ + @{method_def auto} & : & @{text method} \\ + @{method_def force} & : & @{text method} \\ + @{method_def fast} & : & @{text method} \\ + @{method_def slow} & : & @{text method} \\ + @{method_def best} & : & @{text method} \\ + @{method_def fastforce} & : & @{text method} \\ + @{method_def slowsimp} & : & @{text method} \\ + @{method_def bestsimp} & : & @{text method} \\ + @{method_def deepen} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method blast} @{syntax nat}? (@{syntax clamod} * ) + ; + @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) + ; + @@{method force} (@{syntax clasimpmod} * ) + ; + (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) + ; + (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) + (@{syntax clasimpmod} * ) + ; + @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) + ; + @{syntax_def clamod}: + (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs} + ; + @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | + ('cong' | 'split') (() | 'add' | 'del') | + 'iff' (((() | 'add') '?'?) | 'del') | + (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs} + "} + + \begin{description} + + \item @{method blast} is a separate classical tableau prover that + uses the same classical rule declarations as explained before. + + Proof search is coded directly in ML using special data structures. + A successful proof is then reconstructed using regular Isabelle + inferences. It is faster and more powerful than the other classical + reasoning tools, but has major limitations too. + + \begin{itemize} + + \item It does not use the classical wrapper tacticals, such as the + integration with the Simplifier of @{method fastforce}. + + \item It does not perform higher-order unification, as needed by the + rule @{thm [source=false] rangeI} in HOL. There are often + alternatives to such rules, for example @{thm [source=false] + range_eqI}. + + \item Function variables may only be applied to parameters of the + subgoal. (This restriction arises because the prover does not use + higher-order unification.) If other function variables are present + then the prover will fail with the message \texttt{Function Var's + argument not a bound variable}. + + \item Its proof strategy is more general than @{method fast} but can + be slower. If @{method blast} fails or seems to be running forever, + try @{method fast} and the other proof tools described below. + + \end{itemize} + + The optional integer argument specifies a bound for the number of + unsafe steps used in a proof. By default, @{method blast} starts + with a bound of 0 and increases it successively to 20. In contrast, + @{text "(blast lim)"} tries to prove the goal using a search bound + of @{text "lim"}. Sometimes a slow proof using @{method blast} can + be made much faster by supplying the successful search bound to this + proof method instead. + + \item @{method auto} combines classical reasoning with + simplification. It is intended for situations where there are a lot + of mostly trivial subgoals; it proves all the easy ones, leaving the + ones it cannot prove. Occasionally, attempting to prove the hard + ones may take a long time. + + The optional depth arguments in @{text "(auto m n)"} refer to its + builtin classical reasoning procedures: @{text m} (default 4) is for + @{method blast}, which is tried first, and @{text n} (default 2) is + for a slower but more general alternative that also takes wrappers + into account. + + \item @{method force} is intended to prove the first subgoal + completely, using many fancy proof tools and performing a rather + exhaustive search. As a result, proof attempts may take rather long + or diverge easily. + + \item @{method fast}, @{method best}, @{method slow} attempt to + prove the first subgoal using sequent-style reasoning as explained + before. Unlike @{method blast}, they construct proofs directly in + Isabelle. + + There is a difference in search strategy and back-tracking: @{method + fast} uses depth-first search and @{method best} uses best-first + search (guided by a heuristic function: normally the total size of + the proof state). + + Method @{method slow} is like @{method fast}, but conducts a broader + search: it may, when backtracking from a failed proof attempt, undo + even the step of proving a subgoal by assumption. + + \item @{method fastforce}, @{method slowsimp}, @{method bestsimp} + are like @{method fast}, @{method slow}, @{method best}, + respectively, but use the Simplifier as additional wrapper. The name + @{method fastforce}, reflects the behaviour of this popular method + better without requiring an understanding of its implementation. + + \item @{method deepen} works by exhaustive search up to a certain + depth. The start depth is 4 (unless specified explicitly), and the + depth is increased iteratively up to 10. Unsafe rules are modified + to preserve the formula they act on, so that it be used repeatedly. + This method can prove more goals than @{method fast}, but is much + slower, for example if the assumptions have many universal + quantifiers. + + \end{description} + + Any of the above methods support additional modifiers of the context + of classical (and simplifier) rules, but the ones related to the + Simplifier are explicitly prefixed by @{text simp} here. The + semantics of these ad-hoc rule declarations is analogous to the + attributes given before. Facts provided by forward chaining are + inserted into the goal before commencing proof search. +*} + + +subsection {* Semi-automated methods *} + +text {* These proof methods may help in situations when the + fully-automated tools fail. The result is a simpler subgoal that + can be tackled by other means, such as by manual instantiation of + quantifiers. + + \begin{matharray}{rcl} + @{method_def safe} & : & @{text method} \\ + @{method_def clarify} & : & @{text method} \\ + @{method_def clarsimp} & : & @{text method} \\ + \end{matharray} + + @{rail " + (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) + ; + @@{method clarsimp} (@{syntax clasimpmod} * ) + "} + + \begin{description} + + \item @{method safe} repeatedly performs safe steps on all subgoals. + It is deterministic, with at most one outcome. + + \item @{method clarify} performs a series of safe steps without + splitting subgoals; see also @{ML clarify_step_tac}. + + \item @{method clarsimp} acts like @{method clarify}, but also does + simplification. Note that if the Simplifier context includes a + splitter for the premises, the subgoal may still be split. + + \end{description} +*} + + +subsection {* Single-step tactics *} + +text {* + \begin{matharray}{rcl} + @{index_ML safe_step_tac: "Proof.context -> int -> tactic"} \\ + @{index_ML inst_step_tac: "Proof.context -> int -> tactic"} \\ + @{index_ML step_tac: "Proof.context -> int -> tactic"} \\ + @{index_ML slow_step_tac: "Proof.context -> int -> tactic"} \\ + @{index_ML clarify_step_tac: "Proof.context -> int -> tactic"} \\ + \end{matharray} + + These are the primitive tactics behind the (semi)automated proof + methods of the Classical Reasoner. By calling them yourself, you + can execute these procedures one step at a time. + + \begin{description} + + \item @{ML safe_step_tac}~@{text "ctxt i"} performs a safe step on + subgoal @{text i}. The safe wrapper tacticals are applied to a + tactic that may include proof by assumption or Modus Ponens (taking + care not to instantiate unknowns), or substitution. + + \item @{ML inst_step_tac} is like @{ML safe_step_tac}, but allows + unknowns to be instantiated. + + \item @{ML step_tac}~@{text "ctxt i"} is the basic step of the proof + procedure. The unsafe wrapper tacticals are applied to a tactic + that tries @{ML safe_tac}, @{ML inst_step_tac}, or applies an unsafe + rule from the context. + + \item @{ML slow_step_tac} resembles @{ML step_tac}, but allows + backtracking between using safe rules with instantiation (@{ML + inst_step_tac}) and using unsafe rules. The resulting search space + is larger. + + \item @{ML clarify_step_tac}~@{text "ctxt i"} performs a safe step + on subgoal @{text i}. No splitting step is applied; for example, + the subgoal @{text "A \ B"} is left as a conjunction. Proof by + assumption, Modus Ponens, etc., may be performed provided they do + not instantiate unknowns. Assumptions of the form @{text "x = t"} + may be eliminated. The safe wrapper tactical is applied. + + \end{description} +*} + + +section {* Object-logic setup \label{sec:object-logic} *} + +text {* + \begin{matharray}{rcl} + @{command_def "judgment"} & : & @{text "theory \ theory"} \\ + @{method_def atomize} & : & @{text method} \\ + @{attribute_def atomize} & : & @{text attribute} \\ + @{attribute_def rule_format} & : & @{text attribute} \\ + @{attribute_def rulify} & : & @{text attribute} \\ + \end{matharray} + + The very starting point for any Isabelle object-logic is a ``truth + judgment'' that links object-level statements to the meta-logic + (with its minimal language of @{text prop} that covers universal + quantification @{text "\"} and implication @{text "\"}). + + Common object-logics are sufficiently expressive to internalize rule + statements over @{text "\"} and @{text "\"} within their own + language. This is useful in certain situations where a rule needs + to be viewed as an atomic statement from the meta-level perspective, + e.g.\ @{text "\x. x \ A \ P x"} versus @{text "\x \ A. P x"}. + + From the following language elements, only the @{method atomize} + method and @{attribute rule_format} attribute are occasionally + required by end-users, the rest is for those who need to setup their + own object-logic. In the latter case existing formulations of + Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. + + Generic tools may refer to the information provided by object-logic + declarations internally. + + @{rail " + @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? + ; + @@{attribute atomize} ('(' 'full' ')')? + ; + @@{attribute rule_format} ('(' 'noasm' ')')? + "} + + \begin{description} + + \item @{command "judgment"}~@{text "c :: \ (mx)"} declares constant + @{text c} as the truth judgment of the current object-logic. Its + type @{text \} should specify a coercion of the category of + object-level propositions to @{text prop} of the Pure meta-logic; + the mixfix annotation @{text "(mx)"} would typically just link the + object language (internally of syntactic category @{text logic}) + with that of @{text prop}. Only one @{command "judgment"} + declaration may be given in any theory development. + + \item @{method atomize} (as a method) rewrites any non-atomic + premises of a sub-goal, using the meta-level equations declared via + @{attribute atomize} (as an attribute) beforehand. As a result, + heavily nested goals become amenable to fundamental operations such + as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``@{text + "(full)"}'' option here means to turn the whole subgoal into an + object-statement (if possible), including the outermost parameters + and assumptions as well. + + A typical collection of @{attribute atomize} rules for a particular + object-logic would provide an internalization for each of the + connectives of @{text "\"}, @{text "\"}, and @{text "\"}. + Meta-level conjunction should be covered as well (this is + particularly important for locales, see \secref{sec:locale}). + + \item @{attribute rule_format} rewrites a theorem by the equalities + declared as @{attribute rulify} rules in the current object-logic. + By default, the result is fully normalized, including assumptions + and conclusions at any depth. The @{text "(no_asm)"} option + restricts the transformation to the conclusion of a rule. + + In common object-logics (HOL, FOL, ZF), the effect of @{attribute + rule_format} is to replace (bounded) universal quantification + (@{text "\"}) and implication (@{text "\"}) by the corresponding + rule statements over @{text "\"} and @{text "\"}. + + \end{description} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/HOL_Specific.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/HOL_Specific.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,2366 @@ +theory HOL_Specific +imports Base Main "~~/src/HOL/Library/Old_Recdef" +begin + +chapter {* Isabelle/HOL \label{ch:hol} *} + +section {* Higher-Order Logic *} + +text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic + version of Church's Simple Theory of Types. HOL can be best + understood as a simply-typed version of classical set theory. The + logic was first implemented in Gordon's HOL system + \cite{mgordon-hol}. It extends Church's original logic + \cite{church40} by explicit type variables (naive polymorphism) and + a sound axiomatization scheme for new types based on subsets of + existing types. + + Andrews's book \cite{andrews86} is a full description of the + original Church-style higher-order logic, with proofs of correctness + and completeness wrt.\ certain set-theoretic interpretations. The + particular extensions of Gordon-style HOL are explained semantically + in two chapters of the 1993 HOL book \cite{pitts93}. + + Experience with HOL over decades has demonstrated that higher-order + logic is widely applicable in many areas of mathematics and computer + science. In a sense, Higher-Order Logic is simpler than First-Order + Logic, because there are fewer restrictions and special cases. Note + that HOL is \emph{weaker} than FOL with axioms for ZF set theory, + which is traditionally considered the standard foundation of regular + mathematics, but for most applications this does not matter. If you + prefer ML to Lisp, you will probably prefer HOL to ZF. + + \medskip The syntax of HOL follows @{text "\"}-calculus and + functional programming. Function application is curried. To apply + the function @{text f} of type @{text "\\<^sub>1 \ \\<^sub>2 \ \\<^sub>3"} to the + arguments @{text a} and @{text b} in HOL, you simply write @{text "f + a b"} (as in ML or Haskell). There is no ``apply'' operator; the + existing application of the Pure @{text "\"}-calculus is re-used. + Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to + the pair @{text "(a, b)"} (which is notation for @{text "Pair a + b"}). The latter typically introduces extra formal efforts that can + be avoided by currying functions by default. Explicit tuples are as + infrequent in HOL formalizations as in good ML or Haskell programs. + + \medskip Isabelle/HOL has a distinct feel, compared to other + object-logics like Isabelle/ZF. It identifies object-level types + with meta-level types, taking advantage of the default + type-inference mechanism of Isabelle/Pure. HOL fully identifies + object-level functions with meta-level functions, with native + abstraction and application. + + These identifications allow Isabelle to support HOL particularly + nicely, but they also mean that HOL requires some sophistication + from the user. In particular, an understanding of Hindley-Milner + type-inference with type-classes, which are both used extensively in + the standard libraries and applications. Beginners can set + @{attribute show_types} or even @{attribute show_sorts} to get more + explicit information about the result of type-inference. *} + + +section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "inductive"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "coinductive"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def (HOL) mono} & : & @{text attribute} \\ + \end{matharray} + + An \emph{inductive definition} specifies the least predicate or set + @{text R} closed under given rules: applying a rule to elements of + @{text R} yields a result within @{text R}. For example, a + structural operational semantics is an inductive definition of an + evaluation relation. + + Dually, a \emph{coinductive definition} specifies the greatest + predicate or set @{text R} that is consistent with given rules: + every element of @{text R} can be seen as arising by applying a rule + to elements of @{text R}. An important example is using + bisimulation relations to formalise equivalence of processes and + infinite data structures. + + Both inductive and coinductive definitions are based on the + Knaster-Tarski fixed-point theorem for complete lattices. The + collection of introduction rules given by the user determines a + functor on subsets of set-theoretic relations. The required + monotonicity of the recursion scheme is proven as a prerequisite to + the fixed-point definition and the resulting consequences. This + works by pushing inclusion through logical connectives and any other + operator that might be wrapped around recursive occurrences of the + defined relation: there must be a monotonicity theorem of the form + @{text "A \ B \ \ A \ \ B"}, for each premise @{text "\ R t"} in an + introduction rule. The default rule declarations of Isabelle/HOL + already take care of most common situations. + + @{rail " + (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | + @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) + @{syntax target}? \\ + @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? (@'where' clauses)? \\ + (@'monos' @{syntax thmrefs})? + ; + clauses: (@{syntax thmdecl}? @{syntax prop} + '|') + ; + @@{attribute (HOL) mono} (() | 'add' | 'del') + "} + + \begin{description} + + \item @{command (HOL) "inductive"} and @{command (HOL) + "coinductive"} define (co)inductive predicates from the introduction + rules. + + The propositions given as @{text "clauses"} in the @{keyword + "where"} part are either rules of the usual @{text "\/\"} format + (with arbitrary nesting), or equalities using @{text "\"}. The + latter specifies extra-logical abbreviations in the sense of + @{command_ref abbreviation}. Introducing abstract syntax + simultaneously with the actual introduction rules is occasionally + useful for complex specifications. + + The optional @{keyword "for"} part contains a list of parameters of + the (co)inductive predicates that remain fixed throughout the + definition, in contrast to arguments of the relation that may vary + in each occurrence within the given @{text "clauses"}. + + The optional @{keyword "monos"} declaration contains additional + \emph{monotonicity theorems}, which are required for each operator + applied to a recursive set in the introduction rules. + + \item @{command (HOL) "inductive_set"} and @{command (HOL) + "coinductive_set"} are wrappers for to the previous commands for + native HOL predicates. This allows to define (co)inductive sets, + where multiple arguments are simulated via tuples. + + \item @{attribute (HOL) mono} declares monotonicity rules in the + context. These rule are involved in the automated monotonicity + proof of the above inductive and coinductive definitions. + + \end{description} +*} + + +subsection {* Derived rules *} + +text {* A (co)inductive definition of @{text R} provides the following + main theorems: + + \begin{description} + + \item @{text R.intros} is the list of introduction rules as proven + theorems, for the recursive predicates (or sets). The rules are + also available individually, using the names given them in the + theory file; + + \item @{text R.cases} is the case analysis (or elimination) rule; + + \item @{text R.induct} or @{text R.coinduct} is the (co)induction + rule; + + \item @{text R.simps} is the equation unrolling the fixpoint of the + predicate one step. + + \end{description} + + When several predicates @{text "R\<^sub>1, \, R\<^sub>n"} are + defined simultaneously, the list of introduction rules is called + @{text "R\<^sub>1_\_R\<^sub>n.intros"}, the case analysis rules are + called @{text "R\<^sub>1.cases, \, R\<^sub>n.cases"}, and the list + of mutual induction rules is called @{text + "R\<^sub>1_\_R\<^sub>n.inducts"}. +*} + + +subsection {* Monotonicity theorems *} + +text {* The context maintains a default set of theorems that are used + in monotonicity proofs. New rules can be declared via the + @{attribute (HOL) mono} attribute. See the main Isabelle/HOL + sources for some examples. The general format of such monotonicity + theorems is as follows: + + \begin{itemize} + + \item Theorems of the form @{text "A \ B \ \ A \ \ B"}, for proving + monotonicity of inductive definitions whose introduction rules have + premises involving terms such as @{text "\ R t"}. + + \item Monotonicity theorems for logical operators, which are of the + general form @{text "(\ \ \) \ \ (\ \ \) \ \ \ \"}. For example, in + the case of the operator @{text "\"}, the corresponding theorem is + \[ + \infer{@{text "P\<^sub>1 \ P\<^sub>2 \ Q\<^sub>1 \ Q\<^sub>2"}}{@{text "P\<^sub>1 \ Q\<^sub>1"} & @{text "P\<^sub>2 \ Q\<^sub>2"}} + \] + + \item De Morgan style equations for reasoning about the ``polarity'' + of expressions, e.g. + \[ + @{prop "\ \ P \ P"} \qquad\qquad + @{prop "\ (P \ Q) \ \ P \ \ Q"} + \] + + \item Equations for reducing complex operators to more primitive + ones whose monotonicity can easily be proved, e.g. + \[ + @{prop "(P \ Q) \ \ P \ Q"} \qquad\qquad + @{prop "Ball A P \ \x. x \ A \ P x"} + \] + + \end{itemize} +*} + +subsubsection {* Examples *} + +text {* The finite powerset operator can be defined inductively like this: *} + +inductive_set Fin :: "'a set \ 'a set set" for A :: "'a set" +where + empty: "{} \ Fin A" +| insert: "a \ A \ B \ Fin A \ insert a B \ Fin A" + +text {* The accessible part of a relation is defined as follows: *} + +inductive acc :: "('a \ 'a \ bool) \ 'a \ bool" + for r :: "'a \ 'a \ bool" (infix "\" 50) +where acc: "(\y. y \ x \ acc r y) \ acc r x" + +text {* Common logical connectives can be easily characterized as +non-recursive inductive definitions with parameters, but without +arguments. *} + +inductive AND for A B :: bool +where "A \ B \ AND A B" + +inductive OR for A B :: bool +where "A \ OR A B" + | "B \ OR A B" + +inductive EXISTS for B :: "'a \ bool" +where "B a \ EXISTS B" + +text {* Here the @{text "cases"} or @{text "induct"} rules produced by + the @{command inductive} package coincide with the expected + elimination rules for Natural Deduction. Already in the original + article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that + each connective can be characterized by its introductions, and the + elimination can be constructed systematically. *} + + +section {* Recursive functions \label{sec:recursion} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "primrec"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "fun"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def (HOL) "function"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def (HOL) "termination"} & : & @{text "local_theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations + ; + (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts? + @{syntax \"fixes\"} \\ @'where' equations + ; + + equations: (@{syntax thmdecl}? @{syntax prop} + '|') + ; + functionopts: '(' (('sequential' | 'domintros') + ',') ')' + ; + @@{command (HOL) termination} @{syntax term}? + "} + + \begin{description} + + \item @{command (HOL) "primrec"} defines primitive recursive + functions over datatypes (see also @{command_ref (HOL) datatype} and + @{command_ref (HOL) rep_datatype}). The given @{text equations} + specify reduction rules that are produced by instantiating the + generic combinator for primitive recursion that is available for + each datatype. + + Each equation needs to be of the form: + + @{text [display] "f x\<^sub>1 \ x\<^sub>m (C y\<^sub>1 \ y\<^sub>k) z\<^sub>1 \ z\<^sub>n = rhs"} + + such that @{text C} is a datatype constructor, @{text rhs} contains + only the free variables on the left-hand side (or from the context), + and all recursive occurrences of @{text "f"} in @{text "rhs"} are of + the form @{text "f \ y\<^sub>i \"} for some @{text i}. At most one + reduction rule for each constructor can be given. The order does + not matter. For missing constructors, the function is defined to + return a default value, but this equation is made difficult to + access for users. + + The reduction rules are declared as @{attribute simp} by default, + which enables standard proof methods like @{method simp} and + @{method auto} to normalize expressions of @{text "f"} applied to + datatype constructions, by simulating symbolic computation via + rewriting. + + \item @{command (HOL) "function"} defines functions by general + wellfounded recursion. A detailed description with examples can be + found in \cite{isabelle-function}. The function is specified by a + set of (possibly conditional) recursive equations with arbitrary + pattern matching. The command generates proof obligations for the + completeness and the compatibility of patterns. + + The defined function is considered partial, and the resulting + simplification rules (named @{text "f.psimps"}) and induction rule + (named @{text "f.pinduct"}) are guarded by a generated domain + predicate @{text "f_dom"}. The @{command (HOL) "termination"} + command can then be used to establish that the function is total. + + \item @{command (HOL) "fun"} is a shorthand notation for ``@{command + (HOL) "function"}~@{text "(sequential)"}, followed by automated + proof attempts regarding pattern matching and termination. See + \cite{isabelle-function} for further details. + + \item @{command (HOL) "termination"}~@{text f} commences a + termination proof for the previously defined function @{text f}. If + this is omitted, the command refers to the most recent function + definition. After the proof is closed, the recursive equations and + the induction principle is established. + + \end{description} + + Recursive definitions introduced by the @{command (HOL) "function"} + command accommodate reasoning by induction (cf.\ @{method induct}): + rule @{text "f.induct"} refers to a specific induction rule, with + parameters named according to the user-specified equations. Cases + are numbered starting from 1. For @{command (HOL) "primrec"}, the + induction principle coincides with structural recursion on the + datatype where the recursion is carried out. + + The equations provided by these packages may be referred later as + theorem list @{text "f.simps"}, where @{text f} is the (collective) + name of the functions defined. Individual equations may be named + explicitly as well. + + The @{command (HOL) "function"} command accepts the following + options. + + \begin{description} + + \item @{text sequential} enables a preprocessor which disambiguates + overlapping patterns by making them mutually disjoint. Earlier + equations take precedence over later ones. This allows to give the + specification in a format very similar to functional programming. + Note that the resulting simplification and induction rules + correspond to the transformed specification, not the one given + originally. This usually means that each equation given by the user + may result in several theorems. Also note that this automatic + transformation only works for ML-style datatype patterns. + + \item @{text domintros} enables the automated generation of + introduction rules for the domain predicate. While mostly not + needed, they can be helpful in some proofs about partial functions. + + \end{description} +*} + +subsubsection {* Example: evaluation of expressions *} + +text {* Subsequently, we define mutual datatypes for arithmetic and + boolean expressions, and use @{command primrec} for evaluation + functions that follow the same recursive structure. *} + +datatype 'a aexp = + IF "'a bexp" "'a aexp" "'a aexp" + | Sum "'a aexp" "'a aexp" + | Diff "'a aexp" "'a aexp" + | Var 'a + | Num nat +and 'a bexp = + Less "'a aexp" "'a aexp" + | And "'a bexp" "'a bexp" + | Neg "'a bexp" + + +text {* \medskip Evaluation of arithmetic and boolean expressions *} + +primrec evala :: "('a \ nat) \ 'a aexp \ nat" + and evalb :: "('a \ nat) \ 'a bexp \ bool" +where + "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)" +| "evala env (Sum a1 a2) = evala env a1 + evala env a2" +| "evala env (Diff a1 a2) = evala env a1 - evala env a2" +| "evala env (Var v) = env v" +| "evala env (Num n) = n" +| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" +| "evalb env (And b1 b2) = (evalb env b1 \ evalb env b2)" +| "evalb env (Neg b) = (\ evalb env b)" + +text {* Since the value of an expression depends on the value of its + variables, the functions @{const evala} and @{const evalb} take an + additional parameter, an \emph{environment} that maps variables to + their values. + + \medskip Substitution on expressions can be defined similarly. The + mapping @{text f} of type @{typ "'a \ 'a aexp"} given as a + parameter is lifted canonically on the types @{typ "'a aexp"} and + @{typ "'a bexp"}, respectively. +*} + +primrec substa :: "('a \ 'b aexp) \ 'a aexp \ 'b aexp" + and substb :: "('a \ 'b aexp) \ 'a bexp \ 'b bexp" +where + "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)" +| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" +| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" +| "substa f (Var v) = f v" +| "substa f (Num n) = Num n" +| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" +| "substb f (And b1 b2) = And (substb f b1) (substb f b2)" +| "substb f (Neg b) = Neg (substb f b)" + +text {* In textbooks about semantics one often finds substitution + theorems, which express the relationship between substitution and + evaluation. For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove + such a theorem by mutual induction, followed by simplification. +*} + +lemma subst_one: + "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a" + "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b" + by (induct a and b) simp_all + +lemma subst_all: + "evala env (substa s a) = evala (\x. evala env (s x)) a" + "evalb env (substb s b) = evalb (\x. evala env (s x)) b" + by (induct a and b) simp_all + + +subsubsection {* Example: a substitution function for terms *} + +text {* Functions on datatypes with nested recursion are also defined + by mutual primitive recursion. *} + +datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list" + +text {* A substitution function on type @{typ "('a, 'b) term"} can be + defined as follows, by working simultaneously on @{typ "('a, 'b) + term list"}: *} + +primrec subst_term :: "('a \ ('a, 'b) term) \ ('a, 'b) term \ ('a, 'b) term" and + subst_term_list :: "('a \ ('a, 'b) term) \ ('a, 'b) term list \ ('a, 'b) term list" +where + "subst_term f (Var a) = f a" +| "subst_term f (App b ts) = App b (subst_term_list f ts)" +| "subst_term_list f [] = []" +| "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts" + +text {* The recursion scheme follows the structure of the unfolded + definition of type @{typ "('a, 'b) term"}. To prove properties of this + substitution function, mutual induction is needed: +*} + +lemma "subst_term (subst_term f1 \ f2) t = subst_term f1 (subst_term f2 t)" and + "subst_term_list (subst_term f1 \ f2) ts = subst_term_list f1 (subst_term_list f2 ts)" + by (induct t and ts) simp_all + + +subsubsection {* Example: a map function for infinitely branching trees *} + +text {* Defining functions on infinitely branching datatypes by + primitive recursion is just as easy. +*} + +datatype 'a tree = Atom 'a | Branch "nat \ 'a tree" + +primrec map_tree :: "('a \ 'b) \ 'a tree \ 'b tree" +where + "map_tree f (Atom a) = Atom (f a)" +| "map_tree f (Branch ts) = Branch (\x. map_tree f (ts x))" + +text {* Note that all occurrences of functions such as @{text ts} + above must be applied to an argument. In particular, @{term + "map_tree f \ ts"} is not allowed here. *} + +text {* Here is a simple composition lemma for @{term map_tree}: *} + +lemma "map_tree g (map_tree f t) = map_tree (g \ f) t" + by (induct t) simp_all + + +subsection {* Proof methods related to recursive definitions *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) pat_completeness} & : & @{text method} \\ + @{method_def (HOL) relation} & : & @{text method} \\ + @{method_def (HOL) lexicographic_order} & : & @{text method} \\ + @{method_def (HOL) size_change} & : & @{text method} \\ + @{method_def (HOL) induction_schema} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method (HOL) relation} @{syntax term} + ; + @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) + ; + @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) + ; + @@{method (HOL) induction_schema} + ; + orders: ( 'max' | 'min' | 'ms' ) * + "} + + \begin{description} + + \item @{method (HOL) pat_completeness} is a specialized method to + solve goals regarding the completeness of pattern matching, as + required by the @{command (HOL) "function"} package (cf.\ + \cite{isabelle-function}). + + \item @{method (HOL) relation}~@{text R} introduces a termination + proof using the relation @{text R}. The resulting proof state will + contain goals expressing that @{text R} is wellfounded, and that the + arguments of recursive calls decrease with respect to @{text R}. + Usually, this method is used as the initial proof step of manual + termination proofs. + + \item @{method (HOL) "lexicographic_order"} attempts a fully + automated termination proof by searching for a lexicographic + combination of size measures on the arguments of the function. The + method accepts the same arguments as the @{method auto} method, + which it uses internally to prove local descents. The @{syntax + clasimpmod} modifiers are accepted (as for @{method auto}). + + In case of failure, extensive information is printed, which can help + to analyse the situation (cf.\ \cite{isabelle-function}). + + \item @{method (HOL) "size_change"} also works on termination goals, + using a variation of the size-change principle, together with a + graph decomposition technique (see \cite{krauss_phd} for details). + Three kinds of orders are used internally: @{text max}, @{text min}, + and @{text ms} (multiset), which is only available when the theory + @{text Multiset} is loaded. When no order kinds are given, they are + tried in order. The search for a termination proof uses SAT solving + internally. + + For local descent proofs, the @{syntax clasimpmod} modifiers are + accepted (as for @{method auto}). + + \item @{method (HOL) induction_schema} derives user-specified + induction rules from well-founded induction and completeness of + patterns. This factors out some operations that are done internally + by the function package and makes them available separately. See + @{file "~~/src/HOL/ex/Induction_Schema.thy"} for examples. + + \end{description} +*} + + +subsection {* Functions with explicit partiality *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "partial_function"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{command (HOL) partial_function} @{syntax target}? + '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\ + @'where' @{syntax thmdecl}? @{syntax prop} + "} + + \begin{description} + + \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines + recursive functions based on fixpoints in complete partial + orders. No termination proof is required from the user or + constructed internally. Instead, the possibility of non-termination + is modelled explicitly in the result type, which contains an + explicit bottom element. + + Pattern matching and mutual recursion are currently not supported. + Thus, the specification consists of a single function described by a + single recursive equation. + + There are no fixed syntactic restrictions on the body of the + function, but the induced functional must be provably monotonic + wrt.\ the underlying order. The monotonicitity proof is performed + internally, and the definition is rejected when it fails. The proof + can be influenced by declaring hints using the + @{attribute (HOL) partial_function_mono} attribute. + + The mandatory @{text mode} argument specifies the mode of operation + of the command, which directly corresponds to a complete partial + order on the result type. By default, the following modes are + defined: + + \begin{description} + + \item @{text option} defines functions that map into the @{type + option} type. Here, the value @{term None} is used to model a + non-terminating computation. Monotonicity requires that if @{term + None} is returned by a recursive call, then the overall result must + also be @{term None}. This is best achieved through the use of the + monadic operator @{const "Option.bind"}. + + \item @{text tailrec} defines functions with an arbitrary result + type and uses the slightly degenerated partial order where @{term + "undefined"} is the bottom element. Now, monotonicity requires that + if @{term undefined} is returned by a recursive call, then the + overall result must also be @{term undefined}. In practice, this is + only satisfied when each recursive call is a tail call, whose result + is directly returned. Thus, this mode of operation allows the + definition of arbitrary tail-recursive functions. + + \end{description} + + Experienced users may define new modes by instantiating the locale + @{const "partial_function_definitions"} appropriately. + + \item @{attribute (HOL) partial_function_mono} declares rules for + use in the internal monononicity proofs of partial function + definitions. + + \end{description} + +*} + + +subsection {* Old-style recursive function definitions (TFL) *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "recdef"} & : & @{text "theory \ theory)"} \\ + @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \ proof(prove)"} \\ + \end{matharray} + + The old TFL commands @{command (HOL) "recdef"} and @{command (HOL) + "recdef_tc"} for defining recursive are mostly obsolete; @{command + (HOL) "function"} or @{command (HOL) "fun"} should be used instead. + + @{rail " + @@{command (HOL) recdef} ('(' @'permissive' ')')? \\ + @{syntax name} @{syntax term} (@{syntax prop} +) hints? + ; + recdeftc @{syntax thmdecl}? tc + ; + hints: '(' @'hints' ( recdefmod * ) ')' + ; + recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') + (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod} + ; + tc: @{syntax nameref} ('(' @{syntax nat} ')')? + "} + + \begin{description} + + \item @{command (HOL) "recdef"} defines general well-founded + recursive functions (using the TFL package), see also + \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells + TFL to recover from failed proof attempts, returning unfinished + results. The @{text recdef_simp}, @{text recdef_cong}, and @{text + recdef_wf} hints refer to auxiliary rules to be used in the internal + automated proof process of TFL. Additional @{syntax clasimpmod} + declarations may be given to tune the context of the Simplifier + (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ + \secref{sec:classical}). + + \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the + proof for leftover termination condition number @{text i} (default + 1) as generated by a @{command (HOL) "recdef"} definition of + constant @{text c}. + + Note that in most cases, @{command (HOL) "recdef"} is able to finish + its internal proofs without manual intervention. + + \end{description} + + \medskip Hints for @{command (HOL) "recdef"} may be also declared + globally, using the following attributes. + + \begin{matharray}{rcl} + @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\ + @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\ + @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | + @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') + "} +*} + + +section {* Datatypes \label{sec:hol-datatype} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "datatype"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "rep_datatype"} & : & @{text "theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) datatype} (spec + @'and') + ; + @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) + ; + + spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|') + ; + cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? + "} + + \begin{description} + + \item @{command (HOL) "datatype"} defines inductive datatypes in + HOL. + + \item @{command (HOL) "rep_datatype"} represents existing types as + datatypes. + + For foundational reasons, some basic types such as @{typ nat}, @{typ + "'a \ 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are + introduced by more primitive means using @{command_ref typedef}. To + recover the rich infrastructure of @{command datatype} (e.g.\ rules + for @{method cases} and @{method induct} and the primitive recursion + combinators), such types may be represented as actual datatypes + later. This is done by specifying the constructors of the desired + type, and giving a proof of the induction rule, distinctness and + injectivity of constructors. + + For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the + representation of the primitive sum type as fully-featured datatype. + + \end{description} + + The generated rules for @{method induct} and @{method cases} provide + case names according to the given constructors, while parameters are + named after the types (see also \secref{sec:cases-induct}). + + See \cite{isabelle-HOL} for more details on datatypes, but beware of + the old-style theory syntax being used there! Apart from proper + proof methods for case-analysis and induction, there are also + emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) + induct_tac} available, see \secref{sec:hol-induct-tac}; these admit + to refer directly to the internal structure of subgoals (including + internally bound parameters). +*} + + +subsubsection {* Examples *} + +text {* We define a type of finite sequences, with slightly different + names than the existing @{typ "'a list"} that is already in @{theory + Main}: *} + +datatype 'a seq = Empty | Seq 'a "'a seq" + +text {* We can now prove some simple lemma by structural induction: *} + +lemma "Seq x xs \ xs" +proof (induct xs arbitrary: x) + case Empty + txt {* This case can be proved using the simplifier: the freeness + properties of the datatype are already declared as @{attribute + simp} rules. *} + show "Seq x Empty \ Empty" + by simp +next + case (Seq y ys) + txt {* The step case is proved similarly. *} + show "Seq x (Seq y ys) \ Seq y ys" + using `Seq y ys \ ys` by simp +qed + +text {* Here is a more succinct version of the same proof: *} + +lemma "Seq x xs \ xs" + by (induct xs arbitrary: x) simp_all + + +section {* Records \label{sec:hol-record} *} + +text {* + In principle, records merely generalize the concept of tuples, where + components may be addressed by labels instead of just position. The + logical infrastructure of records in Isabelle/HOL is slightly more + advanced, though, supporting truly extensible record schemes. This + admits operations that are polymorphic with respect to record + extension, yielding ``object-oriented'' effects like (single) + inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more + details on object-oriented verification and record subtyping in HOL. +*} + + +subsection {* Basic concepts *} + +text {* + Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records + at the level of terms and types. The notation is as follows: + + \begin{center} + \begin{tabular}{l|l|l} + & record terms & record types \\ \hline + fixed & @{text "\x = a, y = b\"} & @{text "\x :: A, y :: B\"} \\ + schematic & @{text "\x = a, y = b, \ = m\"} & + @{text "\x :: A, y :: B, \ :: M\"} \\ + \end{tabular} + \end{center} + + \noindent The ASCII representation of @{text "\x = a\"} is @{text + "(| x = a |)"}. + + A fixed record @{text "\x = a, y = b\"} has field @{text x} of value + @{text a} and field @{text y} of value @{text b}. The corresponding + type is @{text "\x :: A, y :: B\"}, assuming that @{text "a :: A"} + and @{text "b :: B"}. + + A record scheme like @{text "\x = a, y = b, \ = m\"} contains fields + @{text x} and @{text y} as before, but also possibly further fields + as indicated by the ``@{text "\"}'' notation (which is actually part + of the syntax). The improper field ``@{text "\"}'' of a record + scheme is called the \emph{more part}. Logically it is just a free + variable, which is occasionally referred to as ``row variable'' in + the literature. The more part of a record scheme may be + instantiated by zero or more further components. For example, the + previous scheme may get instantiated to @{text "\x = a, y = b, z = + c, \ = m'\"}, where @{text m'} refers to a different more part. + Fixed records are special instances of record schemes, where + ``@{text "\"}'' is properly terminated by the @{text "() :: unit"} + element. In fact, @{text "\x = a, y = b\"} is just an abbreviation + for @{text "\x = a, y = b, \ = ()\"}. + + \medskip Two key observations make extensible records in a simply + typed language like HOL work out: + + \begin{enumerate} + + \item the more part is internalized, as a free term or type + variable, + + \item field names are externalized, they cannot be accessed within + the logic as first-class values. + + \end{enumerate} + + \medskip In Isabelle/HOL record types have to be defined explicitly, + fixing their field names and types, and their (optional) parent + record. Afterwards, records may be formed using above syntax, while + obeying the canonical order of fields as given by their declaration. + The record package provides several standard operations like + selectors and updates. The common setup for various generic proof + tools enable succinct reasoning patterns. See also the Isabelle/HOL + tutorial \cite{isabelle-hol-book} for further instructions on using + records in practice. +*} + + +subsection {* Record specifications *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "record"} & : & @{text "theory \ theory"} \\ + \end{matharray} + + @{rail " + @@{command (HOL) record} @{syntax typespec_sorts} '=' \\ + (@{syntax type} '+')? (constdecl +) + ; + constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? + "} + + \begin{description} + + \item @{command (HOL) "record"}~@{text "(\\<^sub>1, \, \\<^sub>m) t = \ + c\<^sub>1 :: \\<^sub>1 + \ c\<^sub>n :: \\<^sub>n"} defines extensible record type @{text "(\\<^sub>1, \, \\<^sub>m) t"}, + derived from the optional parent record @{text "\"} by adding new + field components @{text "c\<^sub>i :: \\<^sub>i"} etc. + + The type variables of @{text "\"} and @{text "\\<^sub>i"} need to be + covered by the (distinct) parameters @{text "\\<^sub>1, \, + \\<^sub>m"}. Type constructor @{text t} has to be new, while @{text + \} needs to specify an instance of an existing record type. At + least one new field @{text "c\<^sub>i"} has to be specified. + Basically, field names need to belong to a unique record. This is + not a real restriction in practice, since fields are qualified by + the record name internally. + + The parent record specification @{text \} is optional; if omitted + @{text t} becomes a root record. The hierarchy of all records + declared within a theory context forms a forest structure, i.e.\ a + set of trees starting with a root record each. There is no way to + merge multiple parent records! + + For convenience, @{text "(\\<^sub>1, \, \\<^sub>m) t"} is made a + type abbreviation for the fixed record type @{text "\c\<^sub>1 :: + \\<^sub>1, \, c\<^sub>n :: \\<^sub>n\"}, likewise is @{text + "(\\<^sub>1, \, \\<^sub>m, \) t_scheme"} made an abbreviation for + @{text "\c\<^sub>1 :: \\<^sub>1, \, c\<^sub>n :: \\<^sub>n, \ :: + \\"}. + + \end{description} +*} + + +subsection {* Record operations *} + +text {* + Any record definition of the form presented above produces certain + standard operations. Selectors and updates are provided for any + field, including the improper one ``@{text more}''. There are also + cumulative record constructor functions. To simplify the + presentation below, we assume for now that @{text "(\\<^sub>1, \, + \\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 :: + \\<^sub>1, \, c\<^sub>n :: \\<^sub>n"}. + + \medskip \textbf{Selectors} and \textbf{updates} are available for + any field (including ``@{text more}''): + + \begin{matharray}{lll} + @{text "c\<^sub>i"} & @{text "::"} & @{text "\\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i"} \\ + @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\\<^sub>i \ \\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>c :: \<^vec>\, \ :: \\"} \\ + \end{matharray} + + There is special syntax for application of updates: @{text "r\x := + a\"} abbreviates term @{text "x_update a r"}. Further notation for + repeated updates is also available: @{text "r\x := a\\y := b\\z := + c\"} may be written @{text "r\x := a, y := b, z := c\"}. Note that + because of postfix notation the order of fields shown here is + reverse than in the actual term. Since repeated updates are just + function applications, fields may be freely permuted in @{text "\x + := a, y := b, z := c\"}, as far as logical equality is concerned. + Thus commutativity of independent updates can be proven within the + logic for any two fields, but not as a general theorem. + + \medskip The \textbf{make} operation provides a cumulative record + constructor function: + + \begin{matharray}{lll} + @{text "t.make"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\"} \\ + \end{matharray} + + \medskip We now reconsider the case of non-root records, which are + derived of some parent. In general, the latter may depend on + another parent as well, resulting in a list of \emph{ancestor + records}. Appending the lists of fields of all ancestors results in + a certain field prefix. The record package automatically takes care + of this by lifting operations over this context of ancestor fields. + Assuming that @{text "(\\<^sub>1, \, \\<^sub>m) t"} has ancestor + fields @{text "b\<^sub>1 :: \\<^sub>1, \, b\<^sub>k :: \\<^sub>k"}, + the above record operations will get the following types: + + \medskip + \begin{tabular}{lll} + @{text "c\<^sub>i"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i"} \\ + @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\\<^sub>i \ + \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ + \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\"} \\ + @{text "t.make"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>k \ \\<^sub>1 \ \ \\<^sub>n \ + \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\"} \\ + \end{tabular} + \medskip + + \noindent Some further operations address the extension aspect of a + derived record scheme specifically: @{text "t.fields"} produces a + record fragment consisting of exactly the new fields introduced here + (the result may serve as a more part elsewhere); @{text "t.extend"} + takes a fixed record and adds a given more part; @{text + "t.truncate"} restricts a record scheme to a fixed record. + + \medskip + \begin{tabular}{lll} + @{text "t.fields"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\"} \\ + @{text "t.extend"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\ \ + \ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\"} \\ + @{text "t.truncate"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\"} \\ + \end{tabular} + \medskip + + \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide + for root records. +*} + + +subsection {* Derived rules and proof tools *} + +text {* + The record package proves several results internally, declaring + these facts to appropriate proof tools. This enables users to + reason about record structures quite conveniently. Assume that + @{text t} is a record type as specified above. + + \begin{enumerate} + + \item Standard conversions for selectors or updates applied to + record constructor terms are made part of the default Simplifier + context; thus proofs by reduction of basic operations merely require + the @{method simp} method without further arguments. These rules + are available as @{text "t.simps"}, too. + + \item Selectors applied to updated records are automatically reduced + by an internal simplification procedure, which is also part of the + standard Simplifier setup. + + \item Inject equations of a form analogous to @{prop "(x, y) = (x', + y') \ x = x' \ y = y'"} are declared to the Simplifier and Classical + Reasoner as @{attribute iff} rules. These rules are available as + @{text "t.iffs"}. + + \item The introduction rule for record equality analogous to @{text + "x r = x r' \ y r = y r' \ \ r = r'"} is declared to the Simplifier, + and as the basic rule context as ``@{attribute intro}@{text "?"}''. + The rule is called @{text "t.equality"}. + + \item Representations of arbitrary record expressions as canonical + constructor terms are provided both in @{method cases} and @{method + induct} format (cf.\ the generic proof methods of the same name, + \secref{sec:cases-induct}). Several variations are available, for + fixed records, record schemes, more parts etc. + + The generic proof methods are sufficiently smart to pick the most + sensible rule according to the type of the indicated record + expression: users just need to apply something like ``@{text "(cases + r)"}'' to a certain proof problem. + + \item The derived record operations @{text "t.make"}, @{text + "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not} + treated automatically, but usually need to be expanded by hand, + using the collective fact @{text "t.defs"}. + + \end{enumerate} +*} + + +subsubsection {* Examples *} + +text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *} + + +section {* Adhoc tuples *} + +text {* + \begin{matharray}{rcl} + @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{attribute (HOL) split_format} ('(' 'complete' ')')? + "} + + \begin{description} + + \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes + arguments in function applications to be represented canonically + according to their tuple type structure. + + Note that this operation tends to invent funny names for new local + parameters introduced. + + \end{description} +*} + + +section {* Typedef axiomatization \label{sec:hol-typedef} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "typedef"} & : & @{text "local_theory \ proof(prove)"} \\ + \end{matharray} + + A Gordon/HOL-style type definition is a certain axiom scheme that + identifies a new type with a subset of an existing type. More + precisely, the new type is defined by exhibiting an existing type + @{text \}, a set @{text "A :: \ set"}, and a theorem that proves + @{prop "\x. x \ A"}. Thus @{text A} is a non-empty subset of @{text + \}, and the new type denotes this subset. New functions are + postulated that establish an isomorphism between the new type and + the subset. In general, the type @{text \} may involve type + variables @{text "\\<^sub>1, \, \\<^sub>n"} which means that the type definition + produces a type constructor @{text "(\\<^sub>1, \, \\<^sub>n) t"} depending on + those type arguments. + + The axiomatization can be considered a ``definition'' in the sense + of the particular set-theoretic interpretation of HOL + \cite{pitts93}, where the universe of types is required to be + downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely + new types introduced by @{command "typedef"} stay within the range + of HOL models by construction. Note that @{command_ref + type_synonym} from Isabelle/Pure merely introduces syntactic + abbreviations, without any logical significance. + + @{rail " + @@{command (HOL) typedef} alt_name? abs_type '=' rep_set + ; + + alt_name: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')' + ; + abs_type: @{syntax typespec_sorts} @{syntax mixfix}? + ; + rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? + "} + + \begin{description} + + \item @{command (HOL) "typedef"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = A"} + axiomatizes a type definition in the background theory of the + current context, depending on a non-emptiness result of the set + @{text A} that needs to be proven here. The set @{text A} may + contain type variables @{text "\\<^sub>1, \, \\<^sub>n"} as specified on the LHS, + but no term variables. + + Even though a local theory specification, the newly introduced type + constructor cannot depend on parameters or assumptions of the + context: this is structurally impossible in HOL. In contrast, the + non-emptiness proof may use local assumptions in unusual situations, + which could result in different interpretations in target contexts: + the meaning of the bijection between the representing set @{text A} + and the new type @{text t} may then change in different application + contexts. + + By default, @{command (HOL) "typedef"} defines both a type + constructor @{text t} for the new type, and a term constant @{text + t} for the representing set within the old type. Use the ``@{text + "(open)"}'' option to suppress a separate constant definition + altogether. The injection from type to set is called @{text Rep_t}, + its inverse @{text Abs_t}, unless explicit @{keyword (HOL) + "morphisms"} specification provides alternative names. + + The core axiomatization uses the locale predicate @{const + type_definition} as defined in Isabelle/HOL. Various basic + consequences of that are instantiated accordingly, re-using the + locale facts with names derived from the new type constructor. Thus + the generic @{thm type_definition.Rep} is turned into the specific + @{text "Rep_t"}, for example. + + Theorems @{thm type_definition.Rep}, @{thm + type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse} + provide the most basic characterization as a corresponding + injection/surjection pair (in both directions). The derived rules + @{thm type_definition.Rep_inject} and @{thm + type_definition.Abs_inject} provide a more convenient version of + injectivity, suitable for automated proof tools (e.g.\ in + declarations involving @{attribute simp} or @{attribute iff}). + Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm + type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/ + @{thm type_definition.Abs_induct} provide alternative views on + surjectivity. These rules are already declared as set or type rules + for the generic @{method cases} and @{method induct} methods, + respectively. + + An alternative name for the set definition (and other derived + entities) may be specified in parentheses; the default is to use + @{text t} directly. + + \end{description} + + \begin{warn} + If you introduce a new type axiomatically, i.e.\ via @{command_ref + typedecl} and @{command_ref axiomatization}, the minimum requirement + is that it has a non-empty model, to avoid immediate collapse of the + HOL logic. Moreover, one needs to demonstrate that the + interpretation of such free-form axiomatizations can coexist with + that of the regular @{command_def typedef} scheme, and any extension + that other people might have introduced elsewhere (e.g.\ in HOLCF + \cite{MuellerNvOS99}). + \end{warn} +*} + +subsubsection {* Examples *} + +text {* Type definitions permit the introduction of abstract data + types in a safe way, namely by providing models based on already + existing types. Given some abstract axiomatic description @{text P} + of a type, this involves two steps: + + \begin{enumerate} + + \item Find an appropriate type @{text \} and subset @{text A} which + has the desired properties @{text P}, and make a type definition + based on this representation. + + \item Prove that @{text P} holds for @{text \} by lifting @{text P} + from the representation. + + \end{enumerate} + + You can later forget about the representation and work solely in + terms of the abstract properties @{text P}. + + \medskip The following trivial example pulls a three-element type + into existence within the formal logical environment of HOL. *} + +typedef three = "{(True, True), (True, False), (False, True)}" + by blast + +definition "One = Abs_three (True, True)" +definition "Two = Abs_three (True, False)" +definition "Three = Abs_three (False, True)" + +lemma three_distinct: "One \ Two" "One \ Three" "Two \ Three" + by (simp_all add: One_def Two_def Three_def Abs_three_inject three_def) + +lemma three_cases: + fixes x :: three obtains "x = One" | "x = Two" | "x = Three" + by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject three_def) + +text {* Note that such trivial constructions are better done with + derived specification mechanisms such as @{command datatype}: *} + +datatype three' = One' | Two' | Three' + +text {* This avoids re-doing basic definitions and proofs from the + primitive @{command typedef} above. *} + + +section {* Functorial structure of types *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \ proof(prove)"} + \end{matharray} + + @{rail " + @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term} + ; + "} + + \begin{description} + + \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to + prove and register properties about the functorial structure of type + constructors. These properties then can be used by other packages + to deal with those type constructors in certain type constructions. + Characteristic theorems are noted in the current local theory. By + default, they are prefixed with the base name of the type + constructor, an explicit prefix can be given alternatively. + + The given term @{text "m"} is considered as \emph{mapper} for the + corresponding type constructor and must conform to the following + type pattern: + + \begin{matharray}{lll} + @{text "m"} & @{text "::"} & + @{text "\\<^isub>1 \ \ \\<^isub>k \ (\<^vec>\\<^isub>n) t \ (\<^vec>\\<^isub>n) t"} \\ + \end{matharray} + + \noindent where @{text t} is the type constructor, @{text + "\<^vec>\\<^isub>n"} and @{text "\<^vec>\\<^isub>n"} are distinct + type variables free in the local theory and @{text "\\<^isub>1"}, + \ldots, @{text "\\<^isub>k"} is a subsequence of @{text "\\<^isub>1 \ + \\<^isub>1"}, @{text "\\<^isub>1 \ \\<^isub>1"}, \ldots, + @{text "\\<^isub>n \ \\<^isub>n"}, @{text "\\<^isub>n \ + \\<^isub>n"}. + + \end{description} +*} + + +section {* Transfer package *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) "transfer"} & : & @{text method} \\ + @{method_def (HOL) "transfer'"} & : & @{text method} \\ + @{method_def (HOL) "transfer_prover"} & : & @{text method} \\ + @{attribute_def (HOL) "transfer_rule"} & : & @{text attribute} \\ + @{attribute_def (HOL) "relator_eq"} & : & @{text attribute} \\ + \end{matharray} + + \begin{description} + + \item @{method (HOL) "transfer"} method replaces the current subgoal + with a logically equivalent one that uses different types and + constants. The replacement of types and constants is guided by the + database of transfer rules. Goals are generalized over all free + variables by default; this is necessary for variables whose types + change, but can be overridden for specific variables with e.g. + @{text "transfer fixing: x y z"}. + + \item @{method (HOL) "transfer'"} is a variant of @{method (HOL) + transfer} that allows replacing a subgoal with one that is + logically stronger (rather than equivalent). For example, a + subgoal involving equality on a quotient type could be replaced + with a subgoal involving equality (instead of the corresponding + equivalence relation) on the underlying raw type. + + \item @{method (HOL) "transfer_prover"} method assists with proving + a transfer rule for a new constant, provided the constant is + defined in terms of other constants that already have transfer + rules. It should be applied after unfolding the constant + definitions. + + \item @{attribute (HOL) "transfer_rule"} attribute maintains a + collection of transfer rules, which relate constants at two + different types. Typical transfer rules may relate different type + instances of the same polymorphic constant, or they may relate an + operation on a raw type to a corresponding operation on an + abstract type (quotient or subtype). For example: + + @{text "((A ===> B) ===> list_all2 A ===> list_all2 B) map map"}\\ + @{text "(cr_int ===> cr_int ===> cr_int) (\(x,y) (u,v). (x+u, y+v)) plus"} + + Lemmas involving predicates on relations can also be registered + using the same attribute. For example: + + @{text "bi_unique A \ (list_all2 A ===> op =) distinct distinct"}\\ + @{text "\bi_unique A; bi_unique B\ \ bi_unique (prod_rel A B)"} + + \item @{attribute (HOL) relator_eq} attribute collects identity laws + for relators of various type constructors, e.g. @{text "list_all2 + (op =) = (op =)"}. The @{method (HOL) transfer} method uses these + lemmas to infer transfer rules for non-polymorphic constants on + the fly. + + \end{description} + +*} + + +section {* Lifting package *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "setup_lifting"} & : & @{text "local_theory \ local_theory"}\\ + @{command_def (HOL) "lift_definition"} & : & @{text "local_theory \ proof(prove)"}\\ + @{command_def (HOL) "print_quotmaps"} & : & @{text "context \"}\\ + @{command_def (HOL) "print_quotients"} & : & @{text "context \"}\\ + @{attribute_def (HOL) "quot_map"} & : & @{text attribute} \\ + @{attribute_def (HOL) "invariant_commute"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{command (HOL) setup_lifting} ('(' 'no_abs_code' ')')? \\ + @{syntax thmref} @{syntax thmref}?; + "} + + @{rail " + @@{command (HOL) lift_definition} @{syntax name} '::' @{syntax type} @{syntax mixfix}? \\ + 'is' @{syntax term}; + "} + + \begin{description} + + \item @{command (HOL) "setup_lifting"} Sets up the Lifting package + to work with a user-defined type. The user must provide either a + quotient theorem @{text "Quotient R Abs Rep T"} or a + type_definition theorem @{text "type_definition Rep Abs A"}. The + package configures transfer rules for equality and quantifiers on + the type, and sets up the @{command_def (HOL) "lift_definition"} + command to work with the type. In the case of a quotient theorem, + an optional theorem @{text "reflp R"} can be provided as a second + argument. This allows the package to generate stronger transfer + rules. + + @{command (HOL) "setup_lifting"} is called automatically if a + quotient type is defined by the command @{command (HOL) + "quotient_type"} from the Quotient package. + + If @{command (HOL) "setup_lifting"} is called with a + type_definition theorem, the abstract type implicitly defined by + the theorem is declared as an abstract type in the code + generator. This allows @{command (HOL) "lift_definition"} to + register (generated) code certificate theorems as abstract code + equations in the code generator. The option @{text "no_abs_code"} + of the command @{command (HOL) "setup_lifting"} can turn off that + behavior and causes that code certificate theorems generated by + @{command (HOL) "lift_definition"} are not registred as abstract + code equations. + + \item @{command (HOL) "lift_definition"} @{text "f :: \ is t"} + Defines a new function @{text f} with an abstract type @{text \} + in terms of a corresponding operation @{text t} on a + representation type. The term @{text t} doesn't have to be + necessarily a constant but it can be any term. + + Users must discharge a respectfulness proof obligation when each + constant is defined. For a type copy, i.e. a typedef with @{text + UNIV}, the proof is discharged automatically. The obligation is + presented in a user-friendly, readable form. A respectfulness + theorem in the standard format @{text f.rsp} and a transfer rule + @{text f.tranfer} for the Transfer package are generated by the + package. + + Integration with code_abstype: For typedefs (e.g. subtypes + corresponding to a datatype invariant, such as dlist), @{command + (HOL) "lift_definition"} generates a code certificate theorem + @{text f.rep_eq} and sets up code generation for each constant. + + \item @{command (HOL) "print_quotmaps"} prints stored quotient map + theorems. + + \item @{command (HOL) "print_quotients"} prints stored quotient + theorems. + + \item @{attribute (HOL) quot_map} registers a quotient map + theorem. For examples see @{file + "~~/src/HOL/Library/Quotient_List.thy"} or other Quotient_*.thy + files. + + \item @{attribute (HOL) invariant_commute} registers a theorem which + shows a relationship between the constant @{text + Lifting.invariant} (used for internal encoding of proper subtypes) + and a relator. Such theorems allows the package to hide @{text + Lifting.invariant} from a user in a user-readable form of a + respectfulness theorem. For examples see @{file + "~~/src/HOL/Library/Quotient_List.thy"} or other Quotient_*.thy + files. + + \end{description} +*} + + +section {* Quotient types *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "quotient_type"} & : & @{text "local_theory \ proof(prove)"}\\ + @{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \ proof(prove)"}\\ + @{command_def (HOL) "print_quotmapsQ3"} & : & @{text "context \"}\\ + @{command_def (HOL) "print_quotientsQ3"} & : & @{text "context \"}\\ + @{command_def (HOL) "print_quotconsts"} & : & @{text "context \"}\\ + @{method_def (HOL) "lifting"} & : & @{text method} \\ + @{method_def (HOL) "lifting_setup"} & : & @{text method} \\ + @{method_def (HOL) "descending"} & : & @{text method} \\ + @{method_def (HOL) "descending_setup"} & : & @{text method} \\ + @{method_def (HOL) "partiality_descending"} & : & @{text method} \\ + @{method_def (HOL) "partiality_descending_setup"} & : & @{text method} \\ + @{method_def (HOL) "regularize"} & : & @{text method} \\ + @{method_def (HOL) "injection"} & : & @{text method} \\ + @{method_def (HOL) "cleaning"} & : & @{text method} \\ + @{attribute_def (HOL) "quot_thm"} & : & @{text attribute} \\ + @{attribute_def (HOL) "quot_lifted"} & : & @{text attribute} \\ + @{attribute_def (HOL) "quot_respect"} & : & @{text attribute} \\ + @{attribute_def (HOL) "quot_preserve"} & : & @{text attribute} \\ + \end{matharray} + + The quotient package defines a new quotient type given a raw type + and a partial equivalence relation. It also includes automation for + transporting definitions and theorems. It can automatically produce + definitions and theorems on the quotient type, given the + corresponding constants and facts on the raw type. + + @{rail " + @@{command (HOL) quotient_type} (spec + @'and'); + + spec: @{syntax typespec} @{syntax mixfix}? '=' \\ + @{syntax type} '/' ('partial' ':')? @{syntax term} \\ + (@'morphisms' @{syntax name} @{syntax name})?; + "} + + @{rail " + @@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \\ + @{syntax term} 'is' @{syntax term}; + + constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? + "} + + @{rail " + @@{method (HOL) lifting} @{syntax thmrefs}? + ; + + @@{method (HOL) lifting_setup} @{syntax thmrefs}? + ; + "} + + \begin{description} + + \item @{command (HOL) "quotient_type"} defines quotient types. The + injection from a quotient type to a raw type is called @{text + rep_t}, its inverse @{text abs_t} unless explicit @{keyword (HOL) + "morphisms"} specification provides alternative names. @{command + (HOL) "quotient_type"} requires the user to prove that the relation + is an equivalence relation (predicate @{text equivp}), unless the + user specifies explicitely @{text partial} in which case the + obligation is @{text part_equivp}. A quotient defined with @{text + partial} is weaker in the sense that less things can be proved + automatically. + + \item @{command (HOL) "quotient_definition"} defines a constant on + the quotient type. + + \item @{command (HOL) "print_quotmapsQ3"} prints quotient map + functions. + + \item @{command (HOL) "print_quotientsQ3"} prints quotients. + + \item @{command (HOL) "print_quotconsts"} prints quotient constants. + + \item @{method (HOL) "lifting"} and @{method (HOL) "lifting_setup"} + methods match the current goal with the given raw theorem to be + lifted producing three new subgoals: regularization, injection and + cleaning subgoals. @{method (HOL) "lifting"} tries to apply the + heuristics for automatically solving these three subgoals and + leaves only the subgoals unsolved by the heuristics to the user as + opposed to @{method (HOL) "lifting_setup"} which leaves the three + subgoals unsolved. + + \item @{method (HOL) "descending"} and @{method (HOL) + "descending_setup"} try to guess a raw statement that would lift + to the current subgoal. Such statement is assumed as a new subgoal + and @{method (HOL) "descending"} continues in the same way as + @{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries + to solve the arising regularization, injection and cleaning + subgoals with the analoguous method @{method (HOL) + "descending_setup"} which leaves the four unsolved subgoals. + + \item @{method (HOL) "partiality_descending"} finds the regularized + theorem that would lift to the current subgoal, lifts it and + leaves as a subgoal. This method can be used with partial + equivalence quotients where the non regularized statements would + not be true. @{method (HOL) "partiality_descending_setup"} leaves + the injection and cleaning subgoals unchanged. + + \item @{method (HOL) "regularize"} applies the regularization + heuristics to the current subgoal. + + \item @{method (HOL) "injection"} applies the injection heuristics + to the current goal using the stored quotient respectfulness + theorems. + + \item @{method (HOL) "cleaning"} applies the injection cleaning + heuristics to the current subgoal using the stored quotient + preservation theorems. + + \item @{attribute (HOL) quot_lifted} attribute tries to + automatically transport the theorem to the quotient type. + The attribute uses all the defined quotients types and quotient + constants often producing undesired results or theorems that + cannot be lifted. + + \item @{attribute (HOL) quot_respect} and @{attribute (HOL) + quot_preserve} attributes declare a theorem as a respectfulness + and preservation theorem respectively. These are stored in the + local theory store and used by the @{method (HOL) "injection"} + and @{method (HOL) "cleaning"} methods respectively. + + \item @{attribute (HOL) quot_thm} declares that a certain theorem + is a quotient extension theorem. Quotient extension theorems + allow for quotienting inside container types. Given a polymorphic + type that serves as a container, a map function defined for this + container using @{command (HOL) "enriched_type"} and a relation + map defined for for the container type, the quotient extension + theorem should be @{term "Quotient3 R Abs Rep \ Quotient3 + (rel_map R) (map Abs) (map Rep)"}. Quotient extension theorems + are stored in a database and are used all the steps of lifting + theorems. + + \end{description} +*} + + +section {* Coercive subtyping *} + +text {* + \begin{matharray}{rcl} + @{attribute_def (HOL) coercion} & : & @{text attribute} \\ + @{attribute_def (HOL) coercion_enabled} & : & @{text attribute} \\ + @{attribute_def (HOL) coercion_map} & : & @{text attribute} \\ + \end{matharray} + + Coercive subtyping allows the user to omit explicit type + conversions, also called \emph{coercions}. Type inference will add + them as necessary when parsing a term. See + \cite{traytel-berghofer-nipkow-2011} for details. + + @{rail " + @@{attribute (HOL) coercion} (@{syntax term})? + ; + "} + @{rail " + @@{attribute (HOL) coercion_map} (@{syntax term})? + ; + "} + + \begin{description} + + \item @{attribute (HOL) "coercion"}~@{text "f"} registers a new + coercion function @{text "f :: \\<^isub>1 \ \\<^isub>2"} where @{text "\\<^isub>1"} and + @{text "\\<^isub>2"} are type constructors without arguments. Coercions are + composed by the inference algorithm if needed. Note that the type + inference algorithm is complete only if the registered coercions + form a lattice. + + \item @{attribute (HOL) "coercion_map"}~@{text "map"} registers a + new map function to lift coercions through type constructors. The + function @{text "map"} must conform to the following type pattern + + \begin{matharray}{lll} + @{text "map"} & @{text "::"} & + @{text "f\<^isub>1 \ \ \ f\<^isub>n \ (\\<^isub>1, \, \\<^isub>n) t \ (\\<^isub>1, \, \\<^isub>n) t"} \\ + \end{matharray} + + where @{text "t"} is a type constructor and @{text "f\<^isub>i"} is of type + @{text "\\<^isub>i \ \\<^isub>i"} or @{text "\\<^isub>i \ \\<^isub>i"}. Registering a map function + overwrites any existing map function for this particular type + constructor. + + \item @{attribute (HOL) "coercion_enabled"} enables the coercion + inference algorithm. + + \end{description} +*} + + +section {* Arithmetic proof support *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) arith} & : & @{text method} \\ + @{attribute_def (HOL) arith} & : & @{text attribute} \\ + @{attribute_def (HOL) arith_split} & : & @{text attribute} \\ + \end{matharray} + + \begin{description} + + \item @{method (HOL) arith} decides linear arithmetic problems (on + types @{text nat}, @{text int}, @{text real}). Any current facts + are inserted into the goal before running the procedure. + + \item @{attribute (HOL) arith} declares facts that are supplied to + the arithmetic provers implicitly. + + \item @{attribute (HOL) arith_split} attribute declares case split + rules to be expanded before @{method (HOL) arith} is invoked. + + \end{description} + + Note that a simpler (but faster) arithmetic prover is already + invoked by the Simplifier. +*} + + +section {* Intuitionistic proof search *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) iprover} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method (HOL) iprover} ( @{syntax rulemod} * ) + "} + + \begin{description} + + \item @{method (HOL) iprover} performs intuitionistic proof search, + depending on specifically declared rules from the context, or given + as explicit arguments. Chained facts are inserted into the goal + before commencing proof search. + + Rules need to be classified as @{attribute (Pure) intro}, + @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the + ``@{text "!"}'' indicator refers to ``safe'' rules, which may be + applied aggressively (without considering back-tracking later). + Rules declared with ``@{text "?"}'' are ignored in proof search (the + single-step @{method (Pure) rule} method still observes these). An + explicit weight annotation may be given as well; otherwise the + number of rule premises will be taken into account here. + + \end{description} +*} + + +section {* Model Elimination and Resolution *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) "meson"} & : & @{text method} \\ + @{method_def (HOL) "metis"} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method (HOL) meson} @{syntax thmrefs}? + ; + + @@{method (HOL) metis} + ('(' ('partial_types' | 'full_types' | 'no_types' | @{syntax name}) ')')? + @{syntax thmrefs}? + "} + + \begin{description} + + \item @{method (HOL) meson} implements Loveland's model elimination + procedure \cite{loveland-78}. See @{file + "~~/src/HOL/ex/Meson_Test.thy"} for examples. + + \item @{method (HOL) metis} combines ordered resolution and ordered + paramodulation to find first-order (or mildly higher-order) proofs. + The first optional argument specifies a type encoding; see the + Sledgehammer manual \cite{isabelle-sledgehammer} for details. The + directory @{file "~~/src/HOL/Metis_Examples"} contains several small + theories developed to a large extent using @{method (HOL) metis}. + + \end{description} +*} + + +section {* Coherent Logic *} + +text {* + \begin{matharray}{rcl} + @{method_def (HOL) "coherent"} & : & @{text method} \\ + \end{matharray} + + @{rail " + @@{method (HOL) coherent} @{syntax thmrefs}? + "} + + \begin{description} + + \item @{method (HOL) coherent} solves problems of \emph{Coherent + Logic} \cite{Bezem-Coquand:2005}, which covers applications in + confluence theory, lattice theory and projective geometry. See + @{file "~~/src/HOL/ex/Coherent.thy"} for some examples. + + \end{description} +*} + + +section {* Proving propositions *} + +text {* + In addition to the standard proof methods, a number of diagnosis + tools search for proofs and provide an Isar proof snippet on success. + These tools are available via the following commands. + + \begin{matharray}{rcl} + @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "try0"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \ theory"} + \end{matharray} + + @{rail " + @@{command (HOL) try} + ; + + @@{command (HOL) try0} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ? + @{syntax nat}? + ; + + @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}? + ; + + @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? ) + ; + + args: ( @{syntax name} '=' value + ',' ) + ; + + facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')' + ; + "} % FIXME check args "value" + + \begin{description} + + \item @{command (HOL) "solve_direct"} checks whether the current + subgoals can be solved directly by an existing theorem. Duplicate + lemmas can be detected in this way. + + \item @{command (HOL) "try0"} attempts to prove a subgoal + using a combination of standard proof methods (@{method auto}, + @{method simp}, @{method blast}, etc.). Additional facts supplied + via @{text "simp:"}, @{text "intro:"}, @{text "elim:"}, and @{text + "dest:"} are passed to the appropriate proof methods. + + \item @{command (HOL) "try"} attempts to prove or disprove a subgoal + using a combination of provers and disprovers (@{command (HOL) + "solve_direct"}, @{command (HOL) "quickcheck"}, @{command (HOL) + "try0"}, @{command (HOL) "sledgehammer"}, @{command (HOL) + "nitpick"}). + + \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal + using external automatic provers (resolution provers and SMT + solvers). See the Sledgehammer manual \cite{isabelle-sledgehammer} + for details. + + \item @{command (HOL) "sledgehammer_params"} changes @{command (HOL) + "sledgehammer"} configuration options persistently. + + \end{description} +*} + + +section {* Checking and refuting propositions *} + +text {* + Identifying incorrect propositions usually involves evaluation of + particular assignments and systematic counterexample search. This + is supported by the following commands. + + \begin{matharray}{rcl} + @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "values"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ + @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "refute_params"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "nitpick_params"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "quickcheck_generator"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "find_unused_assms"} & : & @{text "context \"} + \end{matharray} + + @{rail " + @@{command (HOL) value} ( '[' @{syntax name} ']' )? modes? @{syntax term} + ; + + @@{command (HOL) values} modes? @{syntax nat}? @{syntax term} + ; + + (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick}) + ( '[' args ']' )? @{syntax nat}? + ; + + (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} | + @@{command (HOL) nitpick_params}) ( '[' args ']' )? + ; + + @@{command (HOL) quickcheck_generator} @{syntax nameref} \\ + 'operations:' ( @{syntax term} +) + ; + + @@{command (HOL) find_unused_assms} @{syntax name}? + ; + + modes: '(' (@{syntax name} +) ')' + ; + + args: ( @{syntax name} '=' value + ',' ) + ; + "} % FIXME check "value" + + \begin{description} + + \item @{command (HOL) "value"}~@{text t} evaluates and prints a + term; optionally @{text modes} can be specified, which are appended + to the current print mode; see \secref{sec:print-modes}. + Internally, the evaluation is performed by registered evaluators, + which are invoked sequentially until a result is returned. + Alternatively a specific evaluator can be selected using square + brackets; typical evaluators use the current set of code equations + to normalize and include @{text simp} for fully symbolic evaluation + using the simplifier, @{text nbe} for \emph{normalization by + evaluation} and \emph{code} for code generation in SML. + + \item @{command (HOL) "values"}~@{text t} enumerates a set + comprehension by evaluation and prints its values up to the given + number of solutions; optionally @{text modes} can be specified, + which are appended to the current print mode; see + \secref{sec:print-modes}. + + \item @{command (HOL) "quickcheck"} tests the current goal for + counterexamples using a series of assignments for its free + variables; by default the first subgoal is tested, an other can be + selected explicitly using an optional goal index. Assignments can + be chosen exhausting the search space upto a given size, or using a + fixed number of random assignments in the search space, or exploring + the search space symbolically using narrowing. By default, + quickcheck uses exhaustive testing. A number of configuration + options are supported for @{command (HOL) "quickcheck"}, notably: + + \begin{description} + + \item[@{text tester}] specifies which testing approach to apply. + There are three testers, @{text exhaustive}, @{text random}, and + @{text narrowing}. An unknown configuration option is treated as + an argument to tester, making @{text "tester ="} optional. When + multiple testers are given, these are applied in parallel. If no + tester is specified, quickcheck uses the testers that are set + active, i.e., configurations @{attribute + quickcheck_exhaustive_active}, @{attribute + quickcheck_random_active}, @{attribute + quickcheck_narrowing_active} are set to true. + + \item[@{text size}] specifies the maximum size of the search space + for assignment values. + + \item[@{text genuine_only}] sets quickcheck only to return genuine + counterexample, but not potentially spurious counterexamples due + to underspecified functions. + + \item[@{text abort_potential}] sets quickcheck to abort once it + found a potentially spurious counterexample and to not continue + to search for a further genuine counterexample. + For this option to be effective, the @{text genuine_only} option + must be set to false. + + \item[@{text eval}] takes a term or a list of terms and evaluates + these terms under the variable assignment found by quickcheck. + This option is currently only supported by the default + (exhaustive) tester. + + \item[@{text iterations}] sets how many sets of assignments are + generated for each particular size. + + \item[@{text no_assms}] specifies whether assumptions in + structured proofs should be ignored. + + \item[@{text locale}] specifies how to process conjectures in + a locale context, i.e., they can be interpreted or expanded. + The option is a whitespace-separated list of the two words + @{text interpret} and @{text expand}. The list determines the + order they are employed. The default setting is to first use + interpretations and then test the expanded conjecture. + The option is only provided as attribute declaration, but not + as parameter to the command. + + \item[@{text timeout}] sets the time limit in seconds. + + \item[@{text default_type}] sets the type(s) generally used to + instantiate type variables. + + \item[@{text report}] if set quickcheck reports how many tests + fulfilled the preconditions. + + \item[@{text use_subtype}] if set quickcheck automatically lifts + conjectures to registered subtypes if possible, and tests the + lifted conjecture. + + \item[@{text quiet}] if set quickcheck does not output anything + while testing. + + \item[@{text verbose}] if set quickcheck informs about the current + size and cardinality while testing. + + \item[@{text expect}] can be used to check if the user's + expectation was met (@{text no_expectation}, @{text + no_counterexample}, or @{text counterexample}). + + \end{description} + + These option can be given within square brackets. + + \item @{command (HOL) "quickcheck_params"} changes @{command (HOL) + "quickcheck"} configuration options persistently. + + \item @{command (HOL) "quickcheck_generator"} creates random and + exhaustive value generators for a given type and operations. It + generates values by using the operations as if they were + constructors of that type. + + \item @{command (HOL) "refute"} tests the current goal for + counterexamples using a reduction to SAT. The following + configuration options are supported: + + \begin{description} + + \item[@{text minsize}] specifies the minimum size (cardinality) of + the models to search for. + + \item[@{text maxsize}] specifies the maximum size (cardinality) of + the models to search for. Nonpositive values mean @{text "\"}. + + \item[@{text maxvars}] specifies the maximum number of Boolean + variables to use when transforming the term into a propositional + formula. Nonpositive values mean @{text "\"}. + + \item[@{text satsolver}] specifies the SAT solver to use. + + \item[@{text no_assms}] specifies whether assumptions in + structured proofs should be ignored. + + \item[@{text maxtime}] sets the time limit in seconds. + + \item[@{text expect}] can be used to check if the user's + expectation was met (@{text genuine}, @{text potential}, @{text + none}, or @{text unknown}). + + \end{description} + + These option can be given within square brackets. + + \item @{command (HOL) "refute_params"} changes @{command (HOL) + "refute"} configuration options persistently. + + \item @{command (HOL) "nitpick"} tests the current goal for + counterexamples using a reduction to first-order relational + logic. See the Nitpick manual \cite{isabelle-nitpick} for details. + + \item @{command (HOL) "nitpick_params"} changes @{command (HOL) + "nitpick"} configuration options persistently. + + \item @{command (HOL) "find_unused_assms"} finds potentially superfluous + assumptions in theorems using quickcheck. + It takes the theory name to be checked for superfluous assumptions as + optional argument. If not provided, it checks the current theory. + Options to the internal quickcheck invocations can be changed with + common configuration declarations. + + \end{description} +*} + + +section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *} + +text {* + The following tools of Isabelle/HOL support cases analysis and + induction in unstructured tactic scripts; see also + \secref{sec:cases-induct} for proper Isar versions of similar ideas. + + \begin{matharray}{rcl} + @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\ + @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\ + @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \ local_theory"} \\ + \end{matharray} + + @{rail " + @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule? + ; + @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule? + ; + @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))? + ; + @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and') + ; + + rule: 'rule' ':' @{syntax thmref} + "} + + \begin{description} + + \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit + to reason about inductive types. Rules are selected according to + the declarations by the @{attribute cases} and @{attribute induct} + attributes, cf.\ \secref{sec:cases-induct}. The @{command (HOL) + datatype} package already takes care of this. + + These unstructured tactics feature both goal addressing and dynamic + instantiation. Note that named rule cases are \emph{not} provided + as would be by the proper @{method cases} and @{method induct} proof + methods (see \secref{sec:cases-induct}). Unlike the @{method + induct} method, @{method induct_tac} does not handle structured rule + statements, only the compact object-logic conclusion of the subgoal + being addressed. + + \item @{method (HOL) ind_cases} and @{command (HOL) + "inductive_cases"} provide an interface to the internal @{ML_text + mk_cases} operation. Rules are simplified in an unrestricted + forward manner. + + While @{method (HOL) ind_cases} is a proof method to apply the + result immediately as elimination rules, @{command (HOL) + "inductive_cases"} provides case split theorems at the theory level + for later use. The @{keyword "for"} argument of the @{method (HOL) + ind_cases} method allows to specify a list of variables that should + be generalized before applying the resulting rule. + + \end{description} +*} + + +section {* Executable code *} + +text {* For validation purposes, it is often useful to \emph{execute} + specifications. In principle, execution could be simulated by + Isabelle's inference kernel, i.e. by a combination of resolution and + simplification. Unfortunately, this approach is rather inefficient. + A more efficient way of executing specifications is to translate + them into a functional programming language such as ML. + + Isabelle provides a generic framework to support code generation + from executable specifications. Isabelle/HOL instantiates these + mechanisms in a way that is amenable to end-user applications. Code + can be generated for functional programs (including overloading + using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml}, + Haskell \cite{haskell-revised-report} and Scala + \cite{scala-overview-tech-report}. Conceptually, code generation is + split up in three steps: \emph{selection} of code theorems, + \emph{translation} into an abstract executable view and + \emph{serialization} to a specific \emph{target language}. + Inductive specifications can be executed using the predicate + compiler which operates within HOL. See \cite{isabelle-codegen} for + an introduction. + + \begin{matharray}{rcl} + @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def (HOL) code} & : & @{text attribute} \\ + @{command_def (HOL) "code_abort"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_datatype"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def (HOL) code_unfold} & : & @{text attribute} \\ + @{attribute_def (HOL) code_post} & : & @{text attribute} \\ + @{attribute_def (HOL) code_abbrev} & : & @{text attribute} \\ + @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def (HOL) "code_const"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_type"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_class"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_instance"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_reserved"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_monad"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_include"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_modulename"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_reflect"} & : & @{text "theory \ theory"} \\ + @{command_def (HOL) "code_pred"} & : & @{text "theory \ proof(prove)"} + \end{matharray} + + @{rail " + @@{command (HOL) export_code} ( constexpr + ) \\ + ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\ + ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ? + ; + + const: @{syntax term} + ; + + constexpr: ( const | 'name._' | '_' ) + ; + + typeconstructor: @{syntax nameref} + ; + + class: @{syntax nameref} + ; + + target: 'SML' | 'OCaml' | 'Haskell' | 'Scala' + ; + + @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )? + ; + + @@{command (HOL) code_abort} ( const + ) + ; + + @@{command (HOL) code_datatype} ( const + ) + ; + + @@{attribute (HOL) code_unfold} ( 'del' ) ? + ; + + @@{attribute (HOL) code_post} ( 'del' ) ? + ; + + @@{attribute (HOL) code_abbrev} + ; + + @@{command (HOL) code_thms} ( constexpr + ) ? + ; + + @@{command (HOL) code_deps} ( constexpr + ) ? + ; + + @@{command (HOL) code_const} (const + @'and') \\ + ( ( '(' target ( syntax ? + @'and' ) ')' ) + ) + ; + + @@{command (HOL) code_type} (typeconstructor + @'and') \\ + ( ( '(' target ( syntax ? + @'and' ) ')' ) + ) + ; + + @@{command (HOL) code_class} (class + @'and') \\ + ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + ) + ; + + @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\ + ( ( '(' target ( '-' ? + @'and' ) ')' ) + ) + ; + + @@{command (HOL) code_reserved} target ( @{syntax string} + ) + ; + + @@{command (HOL) code_monad} const const target + ; + + @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') ) + ; + + @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + ) + ; + + @@{command (HOL) code_reflect} @{syntax string} \\ + ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\ + ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ? + ; + + @@{command (HOL) code_pred} \\( '(' @'modes' ':' modedecl ')')? \\ const + ; + + syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string} + ; + + modedecl: (modes | ((const ':' modes) \\ + (@'and' ((const ':' modes @'and') +))?)) + ; + + modes: mode @'as' const + "} + + \begin{description} + + \item @{command (HOL) "export_code"} generates code for a given list + of constants in the specified target language(s). If no + serialization instruction is given, only abstract code is generated + internally. + + Constants may be specified by giving them literally, referring to + all executable contants within a certain theory by giving @{text + "name.*"}, or referring to \emph{all} executable constants currently + available by giving @{text "*"}. + + By default, for each involved theory one corresponding name space + module is generated. Alternativly, a module name may be specified + after the @{keyword "module_name"} keyword; then \emph{all} code is + placed in this module. + + For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification + refers to a single file; for \emph{Haskell}, it refers to a whole + directory, where code is generated in multiple files reflecting the + module hierarchy. Omitting the file specification denotes standard + output. + + Serializers take an optional list of arguments in parentheses. For + \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits + explicit module signatures. + + For \emph{Haskell} a module name prefix may be given using the + ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a + ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate + datatype declaration. + + \item @{attribute (HOL) code} explicitly selects (or with option + ``@{text "del"}'' deselects) a code equation for code generation. + Usually packages introducing code equations provide a reasonable + default setup for selection. Variants @{text "code abstype"} and + @{text "code abstract"} declare abstract datatype certificates or + code equations on abstract datatype representations respectively. + + \item @{command (HOL) "code_abort"} declares constants which are not + required to have a definition by means of code equations; if needed + these are implemented by program abort instead. + + \item @{command (HOL) "code_datatype"} specifies a constructor set + for a logical type. + + \item @{command (HOL) "print_codesetup"} gives an overview on + selected code equations and code generator datatypes. + + \item @{attribute (HOL) code_unfold} declares (or with option + ``@{text "del"}'' removes) theorems which during preprocessing + are applied as rewrite rules to any code equation or evaluation + input. + + \item @{attribute (HOL) code_post} declares (or with option ``@{text + "del"}'' removes) theorems which are applied as rewrite rules to any + result of an evaluation. + + \item @{attribute (HOL) code_abbrev} declares equations which are + applied as rewrite rules to any result of an evaluation and + symmetrically during preprocessing to any code equation or evaluation + input. + + \item @{command (HOL) "print_codeproc"} prints the setup of the code + generator preprocessor. + + \item @{command (HOL) "code_thms"} prints a list of theorems + representing the corresponding program containing all given + constants after preprocessing. + + \item @{command (HOL) "code_deps"} visualizes dependencies of + theorems representing the corresponding program containing all given + constants after preprocessing. + + \item @{command (HOL) "code_const"} associates a list of constants + with target-specific serializations; omitting a serialization + deletes an existing serialization. + + \item @{command (HOL) "code_type"} associates a list of type + constructors with target-specific serializations; omitting a + serialization deletes an existing serialization. + + \item @{command (HOL) "code_class"} associates a list of classes + with target-specific class names; omitting a serialization deletes + an existing serialization. This applies only to \emph{Haskell}. + + \item @{command (HOL) "code_instance"} declares a list of type + constructor / class instance relations as ``already present'' for a + given target. Omitting a ``@{text "-"}'' deletes an existing + ``already present'' declaration. This applies only to + \emph{Haskell}. + + \item @{command (HOL) "code_reserved"} declares a list of names as + reserved for a given target, preventing it to be shadowed by any + generated code. + + \item @{command (HOL) "code_monad"} provides an auxiliary mechanism + to generate monadic code for Haskell. + + \item @{command (HOL) "code_include"} adds arbitrary named content + (``include'') to generated code. A ``@{text "-"}'' as last argument + will remove an already added ``include''. + + \item @{command (HOL) "code_modulename"} declares aliasings from one + module name onto another. + + \item @{command (HOL) "code_reflect"} without a ``@{text "file"}'' + argument compiles code into the system runtime environment and + modifies the code generator setup that future invocations of system + runtime code generation referring to one of the ``@{text + "datatypes"}'' or ``@{text "functions"}'' entities use these + precompiled entities. With a ``@{text "file"}'' argument, the + corresponding code is generated into that specified file without + modifying the code generator setup. + + \item @{command (HOL) "code_pred"} creates code equations for a + predicate given a set of introduction rules. Optional mode + annotations determine which arguments are supposed to be input or + output. If alternative introduction rules are declared, one must + prove a corresponding elimination rule. + + \end{description} +*} + + +section {* Definition by specification \label{sec:hol-specification} *} + +text {* + \begin{matharray}{rcl} + @{command_def (HOL) "specification"} & : & @{text "theory \ proof(prove)"} \\ + @{command_def (HOL) "ax_specification"} & : & @{text "theory \ proof(prove)"} \\ + \end{matharray} + + @{rail " + (@@{command (HOL) specification} | @@{command (HOL) ax_specification}) + '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +) + ; + decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?) + "} + + \begin{description} + + \item @{command (HOL) "specification"}~@{text "decls \"} sets up a + goal stating the existence of terms with the properties specified to + hold for the constants given in @{text decls}. After finishing the + proof, the theory will be augmented with definitions for the given + constants, as well as with theorems stating the properties for these + constants. + + \item @{command (HOL) "ax_specification"}~@{text "decls \"} sets up + a goal stating the existence of terms with the properties specified + to hold for the constants given in @{text decls}. After finishing + the proof, the theory will be augmented with axioms expressing the + properties given in the first place. + + \item @{text decl} declares a constant to be defined by the + specification given. The definition for the constant @{text c} is + bound to the name @{text c_def} unless a theorem name is given in + the declaration. Overloaded constants should be declared as such. + + \end{description} + + Whether to use @{command (HOL) "specification"} or @{command (HOL) + "ax_specification"} is to some extent a matter of style. @{command + (HOL) "specification"} introduces no new axioms, and so by + construction cannot introduce inconsistencies, whereas @{command + (HOL) "ax_specification"} does introduce axioms, but only after the + user has explicitly proven it to be safe. A practical issue must be + considered, though: After introducing two constants with the same + properties using @{command (HOL) "specification"}, one can prove + that the two constants are, in fact, equal. If this might be a + problem, one should use @{command (HOL) "ax_specification"}. +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Inner_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Inner_Syntax.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1574 @@ +theory Inner_Syntax +imports Base Main +begin + +chapter {* Inner syntax --- the term language \label{ch:inner-syntax} *} + +text {* The inner syntax of Isabelle provides concrete notation for + the main entities of the logical framework, notably @{text + "\"}-terms with types and type classes. Applications may either + extend existing syntactic categories by additional notation, or + define new sub-languages that are linked to the standard term + language via some explicit markers. For example @{verbatim + FOO}~@{text "foo"} could embed the syntax corresponding for some + user-defined nonterminal @{text "foo"} --- within the bounds of the + given lexical syntax of Isabelle/Pure. + + The most basic way to specify concrete syntax for logical entities + works via mixfix annotations (\secref{sec:mixfix}), which may be + usually given as part of the original declaration or via explicit + notation commands later on (\secref{sec:notation}). This already + covers many needs of concrete syntax without having to understand + the full complexity of inner syntax layers. + + Further details of the syntax engine involves the classical + distinction of lexical language versus context-free grammar (see + \secref{sec:pure-syntax}), and various mechanisms for \emph{syntax + transformations} (see \secref{sec:syntax-transformations}). +*} + + +section {* Printing logical entities *} + +subsection {* Diagnostic commands \label{sec:print-diag} *} + +text {* + \begin{matharray}{rcl} + @{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + \end{matharray} + + These diagnostic commands assist interactive development by printing + internal logical entities in a human-readable fashion. + + @{rail " + @@{command typ} @{syntax modes}? @{syntax type} ('::' @{syntax sort})? + ; + @@{command term} @{syntax modes}? @{syntax term} + ; + @@{command prop} @{syntax modes}? @{syntax prop} + ; + @@{command thm} @{syntax modes}? @{syntax thmrefs} + ; + ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thmrefs}? + ; + @@{command pr} @{syntax modes}? @{syntax nat}? + ; + + @{syntax_def modes}: '(' (@{syntax name} + ) ')' + "} + + \begin{description} + + \item @{command "typ"}~@{text \} reads and prints a type expression + according to the current context. + + \item @{command "typ"}~@{text "\ :: s"} uses type-inference to + determine the most general way to make @{text "\"} conform to sort + @{text "s"}. For concrete @{text "\"} this checks if the type + belongs to that sort. Dummy type parameters ``@{text "_"}'' + (underscore) are assigned to fresh type variables with most general + sorts, according the the principles of type-inference. + + \item @{command "term"}~@{text t} and @{command "prop"}~@{text \} + read, type-check and print terms or propositions according to the + current theory or proof context; the inferred type of @{text t} is + output as well. Note that these commands are also useful in + inspecting the current environment of term abbreviations. + + \item @{command "thm"}~@{text "a\<^sub>1 \ a\<^sub>n"} retrieves + theorems from the current theory or proof context. Note that any + attributes included in the theorem specifications are applied to a + temporary context derived from the current theory or proof; the + result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1, + \, a\<^sub>n"} do not have any permanent effect. + + \item @{command "prf"} displays the (compact) proof term of the + current proof state (if present), or of the given theorems. Note + that this requires proof terms to be switched on for the current + object logic (see the ``Proof terms'' section of the Isabelle + reference manual for information on how to do this). + + \item @{command "full_prf"} is like @{command "prf"}, but displays + the full proof term, i.e.\ also displays information omitted in the + compact proof term, which is denoted by ``@{text _}'' placeholders + there. + + \item @{command "pr"}~@{text "goals"} prints the current proof state + (if present), including current facts and goals. The optional limit + arguments affect the number of goals to be displayed, which is + initially 10. Omitting limit value leaves the current setting + unchanged. + + \end{description} + + All of the diagnostic commands above admit a list of @{text modes} + to be specified, which is appended to the current print mode; see + also \secref{sec:print-modes}. Thus the output behavior may be + modified according particular print mode features. For example, + @{command "pr"}~@{text "(latex xsymbols)"} would print the current + proof state with mathematical symbols and special characters + represented in {\LaTeX} source, according to the Isabelle style + \cite{isabelle-sys}. + + Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more + systematic way to include formal items into the printed text + document. +*} + + +subsection {* Details of printed content *} + +text {* + \begin{tabular}{rcll} + @{attribute_def show_types} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_sorts} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_consts} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_abbrevs} & : & @{text attribute} & default @{text true} \\ + @{attribute_def show_brackets} & : & @{text attribute} & default @{text false} \\ + @{attribute_def names_long} & : & @{text attribute} & default @{text false} \\ + @{attribute_def names_short} & : & @{text attribute} & default @{text false} \\ + @{attribute_def names_unique} & : & @{text attribute} & default @{text true} \\ + @{attribute_def eta_contract} & : & @{text attribute} & default @{text true} \\ + @{attribute_def goals_limit} & : & @{text attribute} & default @{text 10} \\ + @{attribute_def show_main_goal} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_hyps} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_tags} & : & @{text attribute} & default @{text false} \\ + @{attribute_def show_question_marks} & : & @{text attribute} & default @{text true} \\ + \end{tabular} + \medskip + + These configuration options control the detail of information that + is displayed for types, terms, theorems, goals etc. See also + \secref{sec:config}. + + \begin{description} + + \item @{attribute show_types} and @{attribute show_sorts} control + printing of type constraints for term variables, and sort + constraints for type variables. By default, neither of these are + shown in output. If @{attribute show_sorts} is enabled, types are + always shown as well. + + Note that displaying types and sorts may explain why a polymorphic + inference rule fails to resolve with some goal, or why a rewrite + rule does not apply as expected. + + \item @{attribute show_consts} controls printing of types of + constants when displaying a goal state. + + Note that the output can be enormous, because polymorphic constants + often occur at several different type instances. + + \item @{attribute show_abbrevs} controls folding of constant + abbreviations. + + \item @{attribute show_brackets} controls bracketing in pretty + printed output. If enabled, all sub-expressions of the pretty + printing tree will be parenthesized, even if this produces malformed + term syntax! This crude way of showing the internal structure of + pretty printed entities may occasionally help to diagnose problems + with operator priorities, for example. + + \item @{attribute names_long}, @{attribute names_short}, and + @{attribute names_unique} control the way of printing fully + qualified internal names in external form. See also + \secref{sec:antiq} for the document antiquotation options of the + same names. + + \item @{attribute eta_contract} controls @{text "\"}-contracted + printing of terms. + + The @{text \}-contraction law asserts @{prop "(\x. f x) \ f"}, + provided @{text x} is not free in @{text f}. It asserts + \emph{extensionality} of functions: @{prop "f \ g"} if @{prop "f x \ + g x"} for all @{text x}. Higher-order unification frequently puts + terms into a fully @{text \}-expanded form. For example, if @{text + F} has type @{text "(\ \ \) \ \"} then its expanded form is @{term + "\h. F (\x. h x)"}. + + Enabling @{attribute eta_contract} makes Isabelle perform @{text + \}-contractions before printing, so that @{term "\h. F (\x. h x)"} + appears simply as @{text F}. + + Note that the distinction between a term and its @{text \}-expanded + form occasionally matters. While higher-order resolution and + rewriting operate modulo @{text "\\\"}-conversion, some other tools + might look at terms more discretely. + + \item @{attribute goals_limit} controls the maximum number of + subgoals to be shown in goal output. + + \item @{attribute show_main_goal} controls whether the main result + to be proven should be displayed. This information might be + relevant for schematic goals, to inspect the current claim that has + been synthesized so far. + + \item @{attribute show_hyps} controls printing of implicit + hypotheses of local facts. Normally, only those hypotheses are + displayed that are \emph{not} covered by the assumptions of the + current context: this situation indicates a fault in some tool being + used. + + By enabling @{attribute show_hyps}, output of \emph{all} hypotheses + can be enforced, which is occasionally useful for diagnostic + purposes. + + \item @{attribute show_tags} controls printing of extra annotations + within theorems, such as internal position information, or the case + names being attached by the attribute @{attribute case_names}. + + Note that the @{attribute tagged} and @{attribute untagged} + attributes provide low-level access to the collection of tags + associated with a theorem. + + \item @{attribute show_question_marks} controls printing of question + marks for schematic variables, such as @{text ?x}. Only the leading + question mark is affected, the remaining text is unchanged + (including proper markup for schematic variables that might be + relevant for user interfaces). + + \end{description} +*} + + +subsection {* Alternative print modes \label{sec:print-modes} *} + +text {* + \begin{mldecls} + @{index_ML print_mode_value: "unit -> string list"} \\ + @{index_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\ + \end{mldecls} + + The \emph{print mode} facility allows to modify various operations + for printing. Commands like @{command typ}, @{command term}, + @{command thm} (see \secref{sec:print-diag}) take additional print + modes as optional argument. The underlying ML operations are as + follows. + + \begin{description} + + \item @{ML "print_mode_value ()"} yields the list of currently + active print mode names. This should be understood as symbolic + representation of certain individual features for printing (with + precedence from left to right). + + \item @{ML Print_Mode.with_modes}~@{text "modes f x"} evaluates + @{text "f x"} in an execution context where the print mode is + prepended by the given @{text "modes"}. This provides a thread-safe + way to augment print modes. It is also monotonic in the set of mode + names: it retains the default print mode that certain + user-interfaces might have installed for their proper functioning! + + \end{description} + + \begin{warn} + The old global reference @{ML print_mode} should never be used + directly in applications. Its main reason for being publicly + accessible is to support historic versions of Proof~General. + \end{warn} + + \medskip The pretty printer for inner syntax maintains alternative + mixfix productions for any print mode name invented by the user, say + in commands like @{command notation} or @{command abbreviation}. + Mode names can be arbitrary, but the following ones have a specific + meaning by convention: + + \begin{itemize} + + \item @{verbatim "\"\""} (the empty string): default mode; + implicitly active as last element in the list of modes. + + \item @{verbatim input}: dummy print mode that is never active; may + be used to specify notation that is only available for input. + + \item @{verbatim internal} dummy print mode that is never active; + used internally in Isabelle/Pure. + + \item @{verbatim xsymbols}: enable proper mathematical symbols + instead of ASCII art.\footnote{This traditional mode name stems from + the ``X-Symbol'' package for old versions Proof~General with XEmacs, + although that package has been superseded by Unicode in recent + years.} + + \item @{verbatim HTML}: additional mode that is active in HTML + presentation of Isabelle theory sources; allows to provide + alternative output notation. + + \item @{verbatim latex}: additional mode that is active in {\LaTeX} + document preparation of Isabelle theory sources; allows to provide + alternative output notation. + + \end{itemize} +*} + + +subsection {* Printing limits *} + +text {* + \begin{mldecls} + @{index_ML Pretty.margin_default: "int Unsynchronized.ref"} \\ + @{index_ML print_depth: "int -> unit"} \\ + \end{mldecls} + + These ML functions set limits for pretty printed text. + + \begin{description} + + \item @{ML Pretty.margin_default} indicates the global default for + the right margin of the built-in pretty printer, with initial value + 76. Note that user-interfaces typically control margins + automatically when resizing windows, or even bypass the formatting + engine of Isabelle/ML altogether and do it within the front end via + Isabelle/Scala. + + \item @{ML print_depth}~@{text n} limits the printing depth of the + ML toplevel pretty printer; the precise effect depends on the ML + compiler and run-time system. Typically @{text n} should be less + than 10. Bigger values such as 100--1000 are useful for debugging. + + \end{description} +*} + + +section {* Mixfix annotations \label{sec:mixfix} *} + +text {* Mixfix annotations specify concrete \emph{inner syntax} of + Isabelle types and terms. Locally fixed parameters in toplevel + theorem statements, locale and class specifications also admit + mixfix annotations in a fairly uniform manner. A mixfix annotation + describes describes the concrete syntax, the translation to abstract + syntax, and the pretty printing. Special case annotations provide a + simple means of specifying infix operators and binders. + + Isabelle mixfix syntax is inspired by {\OBJ} \cite{OBJ}. It allows + to specify any context-free priority grammar, which is more general + than the fixity declarations of ML and Prolog. + + @{rail " + @{syntax_def mixfix}: '(' mfix ')' + ; + @{syntax_def struct_mixfix}: '(' ( mfix | @'structure' ) ')' + ; + + mfix: @{syntax template} prios? @{syntax nat}? | + (@'infix' | @'infixl' | @'infixr') @{syntax template} @{syntax nat} | + @'binder' @{syntax template} prios? @{syntax nat} + ; + template: string + ; + prios: '[' (@{syntax nat} + ',') ']' + "} + + The string given as @{text template} may include literal text, + spacing, blocks, and arguments (denoted by ``@{text _}''); the + special symbol ``@{verbatim "\"}'' (printed as ``@{text "\"}'') + represents an index argument that specifies an implicit structure + reference (see also \secref{sec:locale}). Infix and binder + declarations provide common abbreviations for particular mixfix + declarations. So in practice, mixfix templates mostly degenerate to + literal text for concrete syntax, such as ``@{verbatim "++"}'' for + an infix symbol. +*} + + +subsection {* The general mixfix form *} + +text {* In full generality, mixfix declarations work as follows. + Suppose a constant @{text "c :: \\<^sub>1 \ \ \\<^sub>n \ \"} is annotated by + @{text "(mixfix [p\<^sub>1, \, p\<^sub>n] p)"}, where @{text "mixfix"} is a string + @{text "d\<^sub>0 _ d\<^sub>1 _ \ _ d\<^sub>n"} consisting of delimiters that surround + argument positions as indicated by underscores. + + Altogether this determines a production for a context-free priority + grammar, where for each argument @{text "i"} the syntactic category + is determined by @{text "\\<^sub>i"} (with priority @{text "p\<^sub>i"}), and the + result category is determined from @{text "\"} (with priority @{text + "p"}). Priority specifications are optional, with default 0 for + arguments and 1000 for the result.\footnote{Omitting priorities is + prone to syntactic ambiguities unless the delimiter tokens determine + fully bracketed notation, as in @{text "if _ then _ else _ fi"}.} + + Since @{text "\"} may be again a function type, the constant + type scheme may have more argument positions than the mixfix + pattern. Printing a nested application @{text "c t\<^sub>1 \ t\<^sub>m"} for + @{text "m > n"} works by attaching concrete notation only to the + innermost part, essentially by printing @{text "(c t\<^sub>1 \ t\<^sub>n) \ t\<^sub>m"} + instead. If a term has fewer arguments than specified in the mixfix + template, the concrete syntax is ignored. + + \medskip A mixfix template may also contain additional directives + for pretty printing, notably spaces, blocks, and breaks. The + general template format is a sequence over any of the following + entities. + + \begin{description} + + \item @{text "d"} is a delimiter, namely a non-empty sequence of + characters other than the following special characters: + + \smallskip + \begin{tabular}{ll} + @{verbatim "'"} & single quote \\ + @{verbatim "_"} & underscore \\ + @{text "\"} & index symbol \\ + @{verbatim "("} & open parenthesis \\ + @{verbatim ")"} & close parenthesis \\ + @{verbatim "/"} & slash \\ + \end{tabular} + \medskip + + \item @{verbatim "'"} escapes the special meaning of these + meta-characters, producing a literal version of the following + character, unless that is a blank. + + A single quote followed by a blank separates delimiters, without + affecting printing, but input tokens may have additional white space + here. + + \item @{verbatim "_"} is an argument position, which stands for a + certain syntactic category in the underlying grammar. + + \item @{text "\"} is an indexed argument position; this is the place + where implicit structure arguments can be attached. + + \item @{text "s"} is a non-empty sequence of spaces for printing. + This and the following specifications do not affect parsing at all. + + \item @{verbatim "("}@{text n} opens a pretty printing block. The + optional number specifies how much indentation to add when a line + break occurs within the block. If the parenthesis is not followed + by digits, the indentation defaults to 0. A block specified via + @{verbatim "(00"} is unbreakable. + + \item @{verbatim ")"} closes a pretty printing block. + + \item @{verbatim "//"} forces a line break. + + \item @{verbatim "/"}@{text s} allows a line break. Here @{text s} + stands for the string of spaces (zero or more) right after the + slash. These spaces are printed if the break is \emph{not} taken. + + \end{description} + + The general idea of pretty printing with blocks and breaks is also + described in \cite{paulson-ml2}; it goes back to \cite{Oppen:1980}. +*} + + +subsection {* Infixes *} + +text {* Infix operators are specified by convenient short forms that + abbreviate general mixfix annotations as follows: + + \begin{center} + \begin{tabular}{lll} + + @{verbatim "("}@{keyword_def "infix"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} + & @{text "\"} & + @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p + 1"}@{verbatim ", "}@{text "p + 1"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ + @{verbatim "("}@{keyword_def "infixl"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} + & @{text "\"} & + @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p"}@{verbatim ", "}@{text "p + 1"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ + @{verbatim "("}@{keyword_def "infixr"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} + & @{text "\"} & + @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p + 1"}@{verbatim ", "}@{text "p"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ + + \end{tabular} + \end{center} + + The mixfix template @{verbatim "\"(_ "}@{text sy}@{verbatim "/ _)\""} + specifies two argument positions; the delimiter is preceded by a + space and followed by a space or line break; the entire phrase is a + pretty printing block. + + The alternative notation @{verbatim "op"}~@{text sy} is introduced + in addition. Thus any infix operator may be written in prefix form + (as in ML), independently of the number of arguments in the term. +*} + + +subsection {* Binders *} + +text {* A \emph{binder} is a variable-binding construct such as a + quantifier. The idea to formalize @{text "\x. b"} as @{text "All + (\x. b)"} for @{text "All :: ('a \ bool) \ bool"} already goes back + to \cite{church40}. Isabelle declarations of certain higher-order + operators may be annotated with @{keyword_def "binder"} annotations + as follows: + + \begin{center} + @{text "c :: "}@{verbatim "\""}@{text "(\\<^sub>1 \ \\<^sub>2) \ \\<^sub>3"}@{verbatim "\" ("}@{keyword "binder"}@{verbatim " \""}@{text "sy"}@{verbatim "\" ["}@{text "p"}@{verbatim "] "}@{text "q"}@{verbatim ")"} + \end{center} + + This introduces concrete binder syntax @{text "sy x. b"}, where + @{text x} is a bound variable of type @{text "\\<^sub>1"}, the body @{text + b} has type @{text "\\<^sub>2"} and the whole term has type @{text "\\<^sub>3"}. + The optional integer @{text p} specifies the syntactic priority of + the body; the default is @{text "q"}, which is also the priority of + the whole construct. + + Internally, the binder syntax is expanded to something like this: + \begin{center} + @{text "c_binder :: "}@{verbatim "\""}@{text "idts \ \\<^sub>2 \ \\<^sub>3"}@{verbatim "\" (\"(3"}@{text sy}@{verbatim "_./ _)\" [0, "}@{text "p"}@{verbatim "] "}@{text "q"}@{verbatim ")"} + \end{center} + + Here @{syntax (inner) idts} is the nonterminal symbol for a list of + identifiers with optional type constraints (see also + \secref{sec:pure-grammar}). The mixfix template @{verbatim + "\"(3"}@{text sy}@{verbatim "_./ _)\""} defines argument positions + for the bound identifiers and the body, separated by a dot with + optional line break; the entire phrase is a pretty printing block of + indentation level 3. Note that there is no extra space after @{text + "sy"}, so it needs to be included user specification if the binder + syntax ends with a token that may be continued by an identifier + token at the start of @{syntax (inner) idts}. + + Furthermore, a syntax translation to transforms @{text "c_binder x\<^sub>1 + \ x\<^sub>n b"} into iterated application @{text "c (\x\<^sub>1. \ c (\x\<^sub>n. b)\)"}. + This works in both directions, for parsing and printing. *} + + +section {* Explicit notation \label{sec:notation} *} + +text {* + \begin{matharray}{rcll} + @{command_def "type_notation"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "no_type_notation"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "notation"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "no_notation"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "write"} & : & @{text "proof(state) \ proof(state)"} \\ + \end{matharray} + + Commands that introduce new logical entities (terms or types) + usually allow to provide mixfix annotations on the spot, which is + convenient for default notation. Nonetheless, the syntax may be + modified later on by declarations for explicit notation. This + allows to add or delete mixfix annotations for of existing logical + entities within the current context. + + @{rail " + (@@{command type_notation} | @@{command no_type_notation}) @{syntax target}? + @{syntax mode}? \\ (@{syntax nameref} @{syntax mixfix} + @'and') + ; + (@@{command notation} | @@{command no_notation}) @{syntax target}? @{syntax mode}? \\ + (@{syntax nameref} @{syntax struct_mixfix} + @'and') + ; + @@{command write} @{syntax mode}? (@{syntax nameref} @{syntax struct_mixfix} + @'and') + "} + + \begin{description} + + \item @{command "type_notation"}~@{text "c (mx)"} associates mixfix + syntax with an existing type constructor. The arity of the + constructor is retrieved from the context. + + \item @{command "no_type_notation"} is similar to @{command + "type_notation"}, but removes the specified syntax annotation from + the present context. + + \item @{command "notation"}~@{text "c (mx)"} associates mixfix + syntax with an existing constant or fixed variable. The type + declaration of the given entity is retrieved from the context. + + \item @{command "no_notation"} is similar to @{command "notation"}, + but removes the specified syntax annotation from the present + context. + + \item @{command "write"} is similar to @{command "notation"}, but + works within an Isar proof body. + + \end{description} +*} + + +section {* The Pure syntax \label{sec:pure-syntax} *} + +subsection {* Lexical matters \label{sec:inner-lex} *} + +text {* The inner lexical syntax vaguely resembles the outer one + (\secref{sec:outer-lex}), but some details are different. There are + two main categories of inner syntax tokens: + + \begin{enumerate} + + \item \emph{delimiters} --- the literal tokens occurring in + productions of the given priority grammar (cf.\ + \secref{sec:priority-grammar}); + + \item \emph{named tokens} --- various categories of identifiers etc. + + \end{enumerate} + + Delimiters override named tokens and may thus render certain + identifiers inaccessible. Sometimes the logical context admits + alternative ways to refer to the same entity, potentially via + qualified names. + + \medskip The categories for named tokens are defined once and for + all as follows, reusing some categories of the outer token syntax + (\secref{sec:outer-lex}). + + \begin{center} + \begin{supertabular}{rcl} + @{syntax_def (inner) id} & = & @{syntax_ref ident} \\ + @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\ + @{syntax_def (inner) var} & = & @{syntax_ref var} \\ + @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\ + @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\ + @{syntax_def (inner) num_token} & = & @{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat} \\ + @{syntax_def (inner) float_token} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ + @{syntax_def (inner) xnum_token} & = & @{verbatim "#"}@{syntax_ref nat}@{text " | "}@{verbatim "#-"}@{syntax_ref nat} \\ + + @{syntax_def (inner) str_token} & = & @{verbatim "''"} @{text "\"} @{verbatim "''"} \\ + \end{supertabular} + \end{center} + + The token categories @{syntax (inner) num_token}, @{syntax (inner) + float_token}, @{syntax (inner) xnum_token}, and @{syntax (inner) + str_token} are not used in Pure. Object-logics may implement numerals + and string constants by adding appropriate syntax declarations, + together with some translation functions (e.g.\ see Isabelle/HOL). + + The derived categories @{syntax_def (inner) num_const}, @{syntax_def + (inner) float_const}, and @{syntax_def (inner) num_const} provide + robust access to the respective tokens: the syntax tree holds a + syntactic constant instead of a free variable. +*} + + +subsection {* Priority grammars \label{sec:priority-grammar} *} + +text {* A context-free grammar consists of a set of \emph{terminal + symbols}, a set of \emph{nonterminal symbols} and a set of + \emph{productions}. Productions have the form @{text "A = \"}, + where @{text A} is a nonterminal and @{text \} is a string of + terminals and nonterminals. One designated nonterminal is called + the \emph{root symbol}. The language defined by the grammar + consists of all strings of terminals that can be derived from the + root symbol by applying productions as rewrite rules. + + The standard Isabelle parser for inner syntax uses a \emph{priority + grammar}. Each nonterminal is decorated by an integer priority: + @{text "A\<^sup>(\<^sup>p\<^sup>)"}. In a derivation, @{text "A\<^sup>(\<^sup>p\<^sup>)"} may be rewritten + using a production @{text "A\<^sup>(\<^sup>q\<^sup>) = \"} only if @{text "p \ q"}. Any + priority grammar can be translated into a normal context-free + grammar by introducing new nonterminals and productions. + + \medskip Formally, a set of context free productions @{text G} + induces a derivation relation @{text "\\<^sub>G"} as follows. Let @{text + \} and @{text \} denote strings of terminal or nonterminal symbols. + Then @{text "\ A\<^sup>(\<^sup>p\<^sup>) \ \\<^sub>G \ \ \"} holds if and only if @{text G} + contains some production @{text "A\<^sup>(\<^sup>q\<^sup>) = \"} for @{text "p \ q"}. + + \medskip The following grammar for arithmetic expressions + demonstrates how binding power and associativity of operators can be + enforced by priorities. + + \begin{center} + \begin{tabular}{rclr} + @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim "("} @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim ")"} \\ + @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim 0} \\ + @{text "A\<^sup>(\<^sup>0\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} \\ + @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} \\ + @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "="} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} \\ + \end{tabular} + \end{center} + The choice of priorities determines that @{verbatim "-"} binds + tighter than @{verbatim "*"}, which binds tighter than @{verbatim + "+"}. Furthermore @{verbatim "+"} associates to the left and + @{verbatim "*"} to the right. + + \medskip For clarity, grammars obey these conventions: + \begin{itemize} + + \item All priorities must lie between 0 and 1000. + + \item Priority 0 on the right-hand side and priority 1000 on the + left-hand side may be omitted. + + \item The production @{text "A\<^sup>(\<^sup>p\<^sup>) = \"} is written as @{text "A = \ + (p)"}, i.e.\ the priority of the left-hand side actually appears in + a column on the far right. + + \item Alternatives are separated by @{text "|"}. + + \item Repetition is indicated by dots @{text "(\)"} in an informal + but obvious way. + + \end{itemize} + + Using these conventions, the example grammar specification above + takes the form: + \begin{center} + \begin{tabular}{rclc} + @{text A} & @{text "="} & @{verbatim "("} @{text A} @{verbatim ")"} \\ + & @{text "|"} & @{verbatim 0} & \qquad\qquad \\ + & @{text "|"} & @{text A} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} & @{text "(0)"} \\ + & @{text "|"} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ + & @{text "|"} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ + \end{tabular} + \end{center} +*} + + +subsection {* The Pure grammar \label{sec:pure-grammar} *} + +text {* The priority grammar of the @{text "Pure"} theory is defined + approximately like this: + + \begin{center} + \begin{supertabular}{rclr} + + @{syntax_def (inner) any} & = & @{text "prop | logic"} \\\\ + + @{syntax_def (inner) prop} & = & @{verbatim "("} @{text prop} @{verbatim ")"} \\ + & @{text "|"} & @{text "prop\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ + & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ + & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{text "\"} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ + & @{text "|"} & @{text "prop\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "&&&"} @{text "prop\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ + & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ + & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{text "\"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ + & @{text "|"} & @{verbatim "[|"} @{text prop} @{verbatim ";"} @{text "\"} @{verbatim ";"} @{text prop} @{verbatim "|]"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ + & @{text "|"} & @{text "\"} @{text prop} @{verbatim ";"} @{text "\"} @{verbatim ";"} @{text prop} @{text "\"} @{text "\"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ + & @{text "|"} & @{verbatim "!!"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ + & @{text "|"} & @{text "\"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ + & @{text "|"} & @{verbatim OFCLASS} @{verbatim "("} @{text type} @{verbatim ","} @{text logic} @{verbatim ")"} \\ + & @{text "|"} & @{verbatim SORT_CONSTRAINT} @{verbatim "("} @{text type} @{verbatim ")"} \\ + & @{text "|"} & @{verbatim TERM} @{text logic} \\ + & @{text "|"} & @{verbatim PROP} @{text aprop} \\\\ + + @{syntax_def (inner) aprop} & = & @{verbatim "("} @{text aprop} @{verbatim ")"} \\ + & @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ + & @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ + & @{text "|"} & @{verbatim XCONST} @{text "id | "}@{verbatim XCONST} @{text "longid"} \\ + & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\\\ + + @{syntax_def (inner) logic} & = & @{verbatim "("} @{text logic} @{verbatim ")"} \\ + & @{text "|"} & @{text "logic\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ + & @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ + & @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ + & @{text "|"} & @{verbatim XCONST} @{text "id | "}@{verbatim XCONST} @{text "longid"} \\ + & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\ + & @{text "|"} & @{text "\ index\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} \\ + & @{text "|"} & @{verbatim "%"} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ + & @{text "|"} & @{text \} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ + & @{text "|"} & @{verbatim op} @{verbatim "=="}@{text " | "}@{verbatim op} @{text "\"}@{text " | "}@{verbatim op} @{verbatim "&&&"} \\ + & @{text "|"} & @{verbatim op} @{verbatim "==>"}@{text " | "}@{verbatim op} @{text "\"} \\ + & @{text "|"} & @{verbatim TYPE} @{verbatim "("} @{text type} @{verbatim ")"} \\\\ + + @{syntax_def (inner) idt} & = & @{verbatim "("} @{text idt} @{verbatim ")"}@{text " | id | "}@{verbatim "_"} \\ + & @{text "|"} & @{text id} @{verbatim "::"} @{text type} & @{text "(0)"} \\ + & @{text "|"} & @{verbatim "_"} @{verbatim "::"} @{text type} & @{text "(0)"} \\\\ + + @{syntax_def (inner) index} & = & @{verbatim "\<^bsub>"} @{text "logic\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "\<^esub>"}@{text " | | \"} \\\\ + + @{syntax_def (inner) idts} & = & @{text "idt | idt\<^sup>(\<^sup>1\<^sup>) idts"} & @{text "(0)"} \\\\ + + @{syntax_def (inner) pttrn} & = & @{text idt} \\\\ + + @{syntax_def (inner) pttrns} & = & @{text "pttrn | pttrn\<^sup>(\<^sup>1\<^sup>) pttrns"} & @{text "(0)"} \\\\ + + @{syntax_def (inner) type} & = & @{verbatim "("} @{text type} @{verbatim ")"} \\ + & @{text "|"} & @{text "tid | tvar | "}@{verbatim "_"} \\ + & @{text "|"} & @{text "tid"} @{verbatim "::"} @{text "sort | tvar "}@{verbatim "::"} @{text "sort | "}@{verbatim "_"} @{verbatim "::"} @{text "sort"} \\ + & @{text "|"} & @{text "type_name | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) type_name"} \\ + & @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim ")"} @{text type_name} \\ + & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ + & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{text "\"} @{text type} & @{text "(0)"} \\ + & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim "]"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ + & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim "]"} @{text "\"} @{text type} & @{text "(0)"} \\ + @{syntax_def (inner) type_name} & = & @{text "id | longid"} \\\\ + + @{syntax_def (inner) sort} & = & @{syntax class_name}~@{text " | "}@{verbatim "{}"} \\ + & @{text "|"} & @{verbatim "{"} @{syntax class_name} @{verbatim ","} @{text "\"} @{verbatim ","} @{syntax class_name} @{verbatim "}"} \\ + @{syntax_def (inner) class_name} & = & @{text "id | longid"} \\ + \end{supertabular} + \end{center} + + \medskip Here literal terminals are printed @{verbatim "verbatim"}; + see also \secref{sec:inner-lex} for further token categories of the + inner syntax. The meaning of the nonterminals defined by the above + grammar is as follows: + + \begin{description} + + \item @{syntax_ref (inner) any} denotes any term. + + \item @{syntax_ref (inner) prop} denotes meta-level propositions, + which are terms of type @{typ prop}. The syntax of such formulae of + the meta-logic is carefully distinguished from usual conventions for + object-logics. In particular, plain @{text "\"}-term notation is + \emph{not} recognized as @{syntax (inner) prop}. + + \item @{syntax_ref (inner) aprop} denotes atomic propositions, which + are embedded into regular @{syntax (inner) prop} by means of an + explicit @{verbatim PROP} token. + + Terms of type @{typ prop} with non-constant head, e.g.\ a plain + variable, are printed in this form. Constants that yield type @{typ + prop} are expected to provide their own concrete syntax; otherwise + the printed version will appear like @{syntax (inner) logic} and + cannot be parsed again as @{syntax (inner) prop}. + + \item @{syntax_ref (inner) logic} denotes arbitrary terms of a + logical type, excluding type @{typ prop}. This is the main + syntactic category of object-logic entities, covering plain @{text + \}-term notation (variables, abstraction, application), plus + anything defined by the user. + + When specifying notation for logical entities, all logical types + (excluding @{typ prop}) are \emph{collapsed} to this single category + of @{syntax (inner) logic}. + + \item @{syntax_ref (inner) index} denotes an optional index term for + indexed syntax. If omitted, it refers to the first @{keyword + "structure"} variable in the context. The special dummy ``@{text + "\"}'' serves as pattern variable in mixfix annotations that + introduce indexed notation. + + \item @{syntax_ref (inner) idt} denotes identifiers, possibly + constrained by types. + + \item @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref + (inner) idt}. This is the most basic category for variables in + iterated binders, such as @{text "\"} or @{text "\"}. + + \item @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns} + denote patterns for abstraction, cases bindings etc. In Pure, these + categories start as a merely copy of @{syntax (inner) idt} and + @{syntax (inner) idts}, respectively. Object-logics may add + additional productions for binding forms. + + \item @{syntax_ref (inner) type} denotes types of the meta-logic. + + \item @{syntax_ref (inner) sort} denotes meta-level sorts. + + \end{description} + + Here are some further explanations of certain syntax features. + + \begin{itemize} + + \item In @{syntax (inner) idts}, note that @{text "x :: nat y"} is + parsed as @{text "x :: (nat y)"}, treating @{text y} like a type + constructor applied to @{text nat}. To avoid this interpretation, + write @{text "(x :: nat) y"} with explicit parentheses. + + \item Similarly, @{text "x :: nat y :: nat"} is parsed as @{text "x :: + (nat y :: nat)"}. The correct form is @{text "(x :: nat) (y :: + nat)"}, or @{text "(x :: nat) y :: nat"} if @{text y} is last in the + sequence of identifiers. + + \item Type constraints for terms bind very weakly. For example, + @{text "x < y :: nat"} is normally parsed as @{text "(x < y) :: + nat"}, unless @{text "<"} has a very low priority, in which case the + input is likely to be ambiguous. The correct form is @{text "x < (y + :: nat)"}. + + \item Constraints may be either written with two literal colons + ``@{verbatim "::"}'' or the double-colon symbol @{verbatim "\"}, + which actually looks exactly the same in some {\LaTeX} styles. + + \item Dummy variables (written as underscore) may occur in different + roles. + + \begin{description} + + \item A type ``@{text "_"}'' or ``@{text "_ :: sort"}'' acts like an + anonymous inference parameter, which is filled-in according to the + most general type produced by the type-checking phase. + + \item A bound ``@{text "_"}'' refers to a vacuous abstraction, where + the body does not refer to the binding introduced here. As in the + term @{term "\x _. x"}, which is @{text "\"}-equivalent to @{text + "\x y. x"}. + + \item A free ``@{text "_"}'' refers to an implicit outer binding. + Higher definitional packages usually allow forms like @{text "f x _ + = x"}. + + \item A schematic ``@{text "_"}'' (within a term pattern, see + \secref{sec:term-decls}) refers to an anonymous variable that is + implicitly abstracted over its context of locally bound variables. + For example, this allows pattern matching of @{text "{x. f x = g + x}"} against @{text "{x. _ = _}"}, or even @{text "{_. _ = _}"} by + using both bound and schematic dummies. + + \end{description} + + \item The three literal dots ``@{verbatim "..."}'' may be also + written as ellipsis symbol @{verbatim "\"}. In both cases this + refers to a special schematic variable, which is bound in the + context. This special term abbreviation works nicely with + calculational reasoning (\secref{sec:calculation}). + + \item @{verbatim CONST} ensures that the given identifier is treated + as constant term, and passed through the parse tree in fully + internalized form. This is particularly relevant for translation + rules (\secref{sec:syn-trans}), notably on the RHS. + + \item @{verbatim XCONST} is similar to @{verbatim CONST}, but + retains the constant name as given. This is only relevant to + translation rules (\secref{sec:syn-trans}), notably on the LHS. + + \end{itemize} +*} + + +subsection {* Inspecting the syntax *} + +text {* + \begin{matharray}{rcl} + @{command_def "print_syntax"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + \end{matharray} + + \begin{description} + + \item @{command "print_syntax"} prints the inner syntax of the + current context. The output can be quite large; the most important + sections are explained below. + + \begin{description} + + \item @{text "lexicon"} lists the delimiters of the inner token + language; see \secref{sec:inner-lex}. + + \item @{text "prods"} lists the productions of the underlying + priority grammar; see \secref{sec:priority-grammar}. + + The nonterminal @{text "A\<^sup>(\<^sup>p\<^sup>)"} is rendered in plain text as @{text + "A[p]"}; delimiters are quoted. Many productions have an extra + @{text "\ => name"}. These names later become the heads of parse + trees; they also guide the pretty printer. + + Productions without such parse tree names are called \emph{copy + productions}. Their right-hand side must have exactly one + nonterminal symbol (or named token). The parser does not create a + new parse tree node for copy productions, but simply returns the + parse tree of the right-hand symbol. + + If the right-hand side of a copy production consists of a single + nonterminal without any delimiters, then it is called a \emph{chain + production}. Chain productions act as abbreviations: conceptually, + they are removed from the grammar by adding new productions. + Priority information attached to chain productions is ignored; only + the dummy value @{text "-1"} is displayed. + + \item @{text "print modes"} lists the alternative print modes + provided by this grammar; see \secref{sec:print-modes}. + + \item @{text "parse_rules"} and @{text "print_rules"} relate to + syntax translations (macros); see \secref{sec:syn-trans}. + + \item @{text "parse_ast_translation"} and @{text + "print_ast_translation"} list sets of constants that invoke + translation functions for abstract syntax trees, which are only + required in very special situations; see \secref{sec:tr-funs}. + + \item @{text "parse_translation"} and @{text "print_translation"} + list the sets of constants that invoke regular translation + functions; see \secref{sec:tr-funs}. + + \end{description} + + \end{description} +*} + + +subsection {* Ambiguity of parsed expressions *} + +text {* + \begin{tabular}{rcll} + @{attribute_def syntax_ambiguity_warning} & : & @{text attribute} & default @{text true} \\ + @{attribute_def syntax_ambiguity_limit} & : & @{text attribute} & default @{text 10} \\ + \end{tabular} + + Depending on the grammar and the given input, parsing may be + ambiguous. Isabelle lets the Earley parser enumerate all possible + parse trees, and then tries to make the best out of the situation. + Terms that cannot be type-checked are filtered out, which often + leads to a unique result in the end. Unlike regular type + reconstruction, which is applied to the whole collection of input + terms simultaneously, the filtering stage only treats each given + term in isolation. Filtering is also not attempted for individual + types or raw ASTs (as required for @{command translations}). + + Certain warning or error messages are printed, depending on the + situation and the given configuration options. Parsing ultimately + fails, if multiple results remain after the filtering phase. + + \begin{description} + + \item @{attribute syntax_ambiguity_warning} controls output of + explicit warning messages about syntax ambiguity. + + \item @{attribute syntax_ambiguity_limit} determines the number of + resulting parse trees that are shown as part of the printed message + in case of an ambiguity. + + \end{description} +*} + + +section {* Syntax transformations \label{sec:syntax-transformations} *} + +text {* The inner syntax engine of Isabelle provides separate + mechanisms to transform parse trees either as rewrite systems on + first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs + or syntactic @{text "\"}-terms (\secref{sec:tr-funs}). This works + both for parsing and printing, as outlined in + \figref{fig:parse-print}. + + \begin{figure}[htbp] + \begin{center} + \begin{tabular}{cl} + string & \\ + @{text "\"} & lexer + parser \\ + parse tree & \\ + @{text "\"} & parse AST translation \\ + AST & \\ + @{text "\"} & AST rewriting (macros) \\ + AST & \\ + @{text "\"} & parse translation \\ + --- pre-term --- & \\ + @{text "\"} & print translation \\ + AST & \\ + @{text "\"} & AST rewriting (macros) \\ + AST & \\ + @{text "\"} & print AST translation \\ + string & + \end{tabular} + \end{center} + \caption{Parsing and printing with translations}\label{fig:parse-print} + \end{figure} + + These intermediate syntax tree formats eventually lead to a pre-term + with all names and binding scopes resolved, but most type + information still missing. Explicit type constraints might be given by + the user, or implicit position information by the system --- both + need to be passed-through carefully by syntax transformations. + + Pre-terms are further processed by the so-called \emph{check} and + \emph{unckeck} phases that are intertwined with type-inference (see + also \cite{isabelle-implementation}). The latter allows to operate + on higher-order abstract syntax with proper binding and type + information already available. + + As a rule of thumb, anything that manipulates bindings of variables + or constants needs to be implemented as syntax transformation (see + below). Anything else is better done via check/uncheck: a prominent + example application is the @{command abbreviation} concept of + Isabelle/Pure. *} + + +subsection {* Abstract syntax trees \label{sec:ast} *} + +text {* The ML datatype @{ML_type Ast.ast} explicitly represents the + intermediate AST format that is used for syntax rewriting + (\secref{sec:syn-trans}). It is defined in ML as follows: + \begin{ttbox} + datatype ast = + Constant of string | + Variable of string | + Appl of ast list + \end{ttbox} + + An AST is either an atom (constant or variable) or a list of (at + least two) subtrees. Occasional diagnostic output of ASTs uses + notation that resembles S-expression of LISP. Constant atoms are + shown as quoted strings, variable atoms as non-quoted strings and + applications as a parenthesized list of subtrees. For example, the + AST + @{ML [display] "Ast.Appl + [Ast.Constant \"_abs\", Ast.Variable \"x\", Ast.Variable \"t\"]"} + is pretty-printed as @{verbatim "(\"_abs\" x t)"}. Note that + @{verbatim "()"} and @{verbatim "(x)"} are excluded as ASTs, because + they have too few subtrees. + + \medskip AST application is merely a pro-forma mechanism to indicate + certain syntactic structures. Thus @{verbatim "(c a b)"} could mean + either term application or type application, depending on the + syntactic context. + + Nested application like @{verbatim "((\"_abs\" x t) u)"} is also + possible, but ASTs are definitely first-order: the syntax constant + @{verbatim "\"_abs\""} does not bind the @{verbatim x} in any way. + Proper bindings are introduced in later stages of the term syntax, + where @{verbatim "(\"_abs\" x t)"} becomes an @{ML Abs} node and + occurrences of @{verbatim x} in @{verbatim t} are replaced by bound + variables (represented as de-Bruijn indices). +*} + + +subsubsection {* AST constants versus variables *} + +text {* Depending on the situation --- input syntax, output syntax, + translation patterns --- the distinction of atomic asts as @{ML + Ast.Constant} versus @{ML Ast.Variable} serves slightly different + purposes. + + Input syntax of a term such as @{text "f a b = c"} does not yet + indicate the scopes of atomic entities @{text "f, a, b, c"}: they + could be global constants or local variables, even bound ones + depending on the context of the term. @{ML Ast.Variable} leaves + this choice still open: later syntax layers (or translation + functions) may capture such a variable to determine its role + specifically, to make it a constant, bound variable, free variable + etc. In contrast, syntax translations that introduce already known + constants would rather do it via @{ML Ast.Constant} to prevent + accidental re-interpretation later on. + + Output syntax turns term constants into @{ML Ast.Constant} and + variables (free or schematic) into @{ML Ast.Variable}. This + information is precise when printing fully formal @{text "\"}-terms. + + In AST translation patterns (\secref{sec:syn-trans}) the system + guesses from the current theory context which atoms should be + treated as constant versus variable for the matching process. + Sometimes this needs to be indicated more explicitly using @{text + "CONST c"} inside the term language. It is also possible to use + @{command syntax} declarations (without mixfix annotation) to + enforce that certain unqualified names are always treated as + constant within the syntax machinery. + + \medskip For ASTs that represent the language of types or sorts, the + situation is much simpler, since the concrete syntax already + distinguishes type variables from type constants (constructors). So + @{text "('a, 'b) foo"} corresponds to an AST application of some + constant for @{text foo} and variable arguments for @{text "'a"} and + @{text "'b"}. Note that the postfix application is merely a feature + of the concrete syntax, while in the AST the constructor occurs in + head position. *} + + +subsubsection {* Authentic syntax names *} + +text {* Naming constant entities within ASTs is another delicate + issue. Unqualified names are looked up in the name space tables in + the last stage of parsing, after all translations have been applied. + Since syntax transformations do not know about this later name + resolution yet, there can be surprises in boundary cases. + + \emph{Authentic syntax names} for @{ML Ast.Constant} avoid this + problem: the fully-qualified constant name with a special prefix for + its formal category (@{text "class"}, @{text "type"}, @{text + "const"}, @{text "fixed"}) represents the information faithfully + within the untyped AST format. Accidental overlap with free or + bound variables is excluded as well. Authentic syntax names work + implicitly in the following situations: + + \begin{itemize} + + \item Input of term constants (or fixed variables) that are + introduced by concrete syntax via @{command notation}: the + correspondence of a particular grammar production to some known term + entity is preserved. + + \item Input of type constants (constructors) and type classes --- + thanks to explicit syntactic distinction independently on the + context. + + \item Output of term constants, type constants, type classes --- + this information is already available from the internal term to be + printed. + + \end{itemize} + + In other words, syntax transformations that operate on input terms + written as prefix applications are difficult to make robust. + Luckily, this case rarely occurs in practice, because syntax forms + to be translated usually correspond to some bits of concrete + notation. *} + + +subsection {* Raw syntax and translations \label{sec:syn-trans} *} + +text {* + \begin{tabular}{rcll} + @{command_def "nonterminal"} & : & @{text "theory \ theory"} \\ + @{command_def "syntax"} & : & @{text "theory \ theory"} \\ + @{command_def "no_syntax"} & : & @{text "theory \ theory"} \\ + @{command_def "translations"} & : & @{text "theory \ theory"} \\ + @{command_def "no_translations"} & : & @{text "theory \ theory"} \\ + @{attribute_def syntax_ast_trace} & : & @{text attribute} & default @{text false} \\ + @{attribute_def syntax_ast_stats} & : & @{text attribute} & default @{text false} \\ + \end{tabular} + + Unlike mixfix notation for existing formal entities + (\secref{sec:notation}), raw syntax declarations provide full access + to the priority grammar of the inner syntax, without any sanity + checks. This includes additional syntactic categories (via + @{command nonterminal}) and free-form grammar productions (via + @{command syntax}). Additional syntax translations (or macros, via + @{command translations}) are required to turn resulting parse trees + into proper representations of formal entities again. + + @{rail " + @@{command nonterminal} (@{syntax name} + @'and') + ; + (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +) + ; + (@@{command translations} | @@{command no_translations}) + (transpat ('==' | '=>' | '<=' | '\' | '\' | '\') transpat +) + ; + + constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? + ; + mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')') + ; + transpat: ('(' @{syntax nameref} ')')? @{syntax string} + "} + + \begin{description} + + \item @{command "nonterminal"}~@{text c} declares a type + constructor @{text c} (without arguments) to act as purely syntactic + type: a nonterminal symbol of the inner syntax. + + \item @{command "syntax"}~@{text "(mode) c :: \ (mx)"} augments the + priority grammar and the pretty printer table for the given print + mode (default @{verbatim "\"\""}). An optional keyword @{keyword_ref + "output"} means that only the pretty printer table is affected. + + Following \secref{sec:mixfix}, the mixfix annotation @{text "mx = + template ps q"} together with type @{text "\ = \\<^sub>1 \ \ \\<^sub>n \ \"} and + specify a grammar production. The @{text template} contains + delimiter tokens that surround @{text "n"} argument positions + (@{verbatim "_"}). The latter correspond to nonterminal symbols + @{text "A\<^sub>i"} derived from the argument types @{text "\\<^sub>i"} as + follows: + \begin{itemize} + + \item @{text "prop"} if @{text "\\<^sub>i = prop"} + + \item @{text "logic"} if @{text "\\<^sub>i = (\)\"} for logical type + constructor @{text "\ \ prop"} + + \item @{text any} if @{text "\\<^sub>i = \"} for type variables + + \item @{text "\"} if @{text "\\<^sub>i = \"} for nonterminal @{text "\"} + (syntactic type constructor) + + \end{itemize} + + Each @{text "A\<^sub>i"} is decorated by priority @{text "p\<^sub>i"} from the + given list @{text "ps"}; misssing priorities default to 0. + + The resulting nonterminal of the production is determined similarly + from type @{text "\"}, with priority @{text "q"} and default 1000. + + \medskip Parsing via this production produces parse trees @{text + "t\<^sub>1, \, t\<^sub>n"} for the argument slots. The resulting parse tree is + composed as @{text "c t\<^sub>1 \ t\<^sub>n"}, by using the syntax constant @{text + "c"} of the syntax declaration. + + Such syntactic constants are invented on the spot, without formal + check wrt.\ existing declarations. It is conventional to use plain + identifiers prefixed by a single underscore (e.g.\ @{text + "_foobar"}). Names should be chosen with care, to avoid clashes + with other syntax declarations. + + \medskip The special case of copy production is specified by @{text + "c = "}@{verbatim "\"\""} (empty string). It means that the + resulting parse tree @{text "t"} is copied directly, without any + further decoration. + + \item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar + declarations (and translations) resulting from @{text decls}, which + are interpreted in the same manner as for @{command "syntax"} above. + + \item @{command "translations"}~@{text rules} specifies syntactic + translation rules (i.e.\ macros) as first-order rewrite rules on + ASTs (\secref{sec:ast}). The theory context maintains two + independent lists translation rules: parse rules (@{verbatim "=>"} + or @{text "\"}) and print rules (@{verbatim "<="} or @{text "\"}). + For convenience, both can be specified simultaneously as parse~/ + print rules (@{verbatim "=="} or @{text "\"}). + + Translation patterns may be prefixed by the syntactic category to be + used for parsing; the default is @{text logic} which means that + regular term syntax is used. Both sides of the syntax translation + rule undergo parsing and parse AST translations + \secref{sec:tr-funs}, in order to perform some fundamental + normalization like @{text "\x y. b \ \x. \y. b"}, but other AST + translation rules are \emph{not} applied recursively here. + + When processing AST patterns, the inner syntax lexer runs in a + different mode that allows identifiers to start with underscore. + This accommodates the usual naming convention for auxiliary syntax + constants --- those that do not have a logical counter part --- by + allowing to specify arbitrary AST applications within the term + syntax, independently of the corresponding concrete syntax. + + Atomic ASTs are distinguished as @{ML Ast.Constant} versus @{ML + Ast.Variable} as follows: a qualified name or syntax constant + declared via @{command syntax}, or parse tree head of concrete + notation becomes @{ML Ast.Constant}, anything else @{ML + Ast.Variable}. Note that @{text CONST} and @{text XCONST} within + the term language (\secref{sec:pure-grammar}) allow to enforce + treatment as constants. + + AST rewrite rules @{text "(lhs, rhs)"} need to obey the following + side-conditions: + + \begin{itemize} + + \item Rules must be left linear: @{text "lhs"} must not contain + repeated variables.\footnote{The deeper reason for this is that AST + equality is not well-defined: different occurrences of the ``same'' + AST could be decorated differently by accidental type-constraints or + source position information, for example.} + + \item Every variable in @{text "rhs"} must also occur in @{text + "lhs"}. + + \end{itemize} + + \item @{command "no_translations"}~@{text rules} removes syntactic + translation rules, which are interpreted in the same manner as for + @{command "translations"} above. + + \item @{attribute syntax_ast_trace} and @{attribute + syntax_ast_stats} control diagnostic output in the AST normalization + process, when translation rules are applied to concrete input or + output. + + \end{description} + + Raw syntax and translations provides a slightly more low-level + access to the grammar and the form of resulting parse trees. It is + often possible to avoid this untyped macro mechanism, and use + type-safe @{command abbreviation} or @{command notation} instead. + Some important situations where @{command syntax} and @{command + translations} are really need are as follows: + + \begin{itemize} + + \item Iterated replacement via recursive @{command translations}. + For example, consider list enumeration @{term "[a, b, c, d]"} as + defined in theory @{theory List} in Isabelle/HOL. + + \item Change of binding status of variables: anything beyond the + built-in @{keyword "binder"} mixfix annotation requires explicit + syntax translations. For example, consider list filter + comprehension @{term "[x \ xs . P]"} as defined in theory @{theory + List} in Isabelle/HOL. + + \end{itemize} +*} + +subsubsection {* Applying translation rules *} + +text {* As a term is being parsed or printed, an AST is generated as + an intermediate form according to \figref{fig:parse-print}. The AST + is normalized by applying translation rules in the manner of a + first-order term rewriting system. We first examine how a single + rule is applied. + + Let @{text "t"} be the abstract syntax tree to be normalized and + @{text "(lhs, rhs)"} some translation rule. A subtree @{text "u"} + of @{text "t"} is called \emph{redex} if it is an instance of @{text + "lhs"}; in this case the pattern @{text "lhs"} is said to match the + object @{text "u"}. A redex matched by @{text "lhs"} may be + replaced by the corresponding instance of @{text "rhs"}, thus + \emph{rewriting} the AST @{text "t"}. Matching requires some notion + of \emph{place-holders} in rule patterns: @{ML Ast.Variable} serves + this purpose. + + More precisely, the matching of the object @{text "u"} against the + pattern @{text "lhs"} is performed as follows: + + \begin{itemize} + + \item Objects of the form @{ML Ast.Variable}~@{text "x"} or @{ML + Ast.Constant}~@{text "x"} are matched by pattern @{ML + Ast.Constant}~@{text "x"}. Thus all atomic ASTs in the object are + treated as (potential) constants, and a successful match makes them + actual constants even before name space resolution (see also + \secref{sec:ast}). + + \item Object @{text "u"} is matched by pattern @{ML + Ast.Variable}~@{text "x"}, binding @{text "x"} to @{text "u"}. + + \item Object @{ML Ast.Appl}~@{text "us"} is matched by @{ML + Ast.Appl}~@{text "ts"} if @{text "us"} and @{text "ts"} have the + same length and each corresponding subtree matches. + + \item In every other case, matching fails. + + \end{itemize} + + A successful match yields a substitution that is applied to @{text + "rhs"}, generating the instance that replaces @{text "u"}. + + Normalizing an AST involves repeatedly applying translation rules + until none are applicable. This works yoyo-like: top-down, + bottom-up, top-down, etc. At each subtree position, rules are + chosen in order of appearance in the theory definitions. + + The configuration options @{attribute syntax_ast_trace} and + @{attribute syntax_ast_stats} might help to understand this process + and diagnose problems. + + \begin{warn} + If syntax translation rules work incorrectly, the output of + @{command_ref print_syntax} with its \emph{rules} sections reveals the + actual internal forms of AST pattern, without potentially confusing + concrete syntax. Recall that AST constants appear as quoted strings + and variables without quotes. + \end{warn} + + \begin{warn} + If @{attribute_ref eta_contract} is set to @{text "true"}, terms + will be @{text "\"}-contracted \emph{before} the AST rewriter sees + them. Thus some abstraction nodes needed for print rules to match + may vanish. For example, @{text "Ball A (\x. P x)"} would contract + to @{text "Ball A P"} and the standard print rule would fail to + apply. This problem can be avoided by hand-written ML translation + functions (see also \secref{sec:tr-funs}), which is in fact the same + mechanism used in built-in @{keyword "binder"} declarations. + \end{warn} +*} + + +subsection {* Syntax translation functions \label{sec:tr-funs} *} + +text {* + \begin{matharray}{rcl} + @{command_def "parse_ast_translation"} & : & @{text "theory \ theory"} \\ + @{command_def "parse_translation"} & : & @{text "theory \ theory"} \\ + @{command_def "print_translation"} & : & @{text "theory \ theory"} \\ + @{command_def "typed_print_translation"} & : & @{text "theory \ theory"} \\ + @{command_def "print_ast_translation"} & : & @{text "theory \ theory"} \\ + @{ML_antiquotation_def "class_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "type_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "const_syntax"} & : & @{text ML_antiquotation} \\ + @{ML_antiquotation_def "syntax_const"} & : & @{text ML_antiquotation} \\ + \end{matharray} + + Syntax translation functions written in ML admit almost arbitrary + manipulations of inner syntax, at the expense of some complexity and + obscurity in the implementation. + + @{rail " + ( @@{command parse_ast_translation} | @@{command parse_translation} | + @@{command print_translation} | @@{command typed_print_translation} | + @@{command print_ast_translation}) ('(' @'advanced' ')')? @{syntax text} + ; + (@@{ML_antiquotation class_syntax} | + @@{ML_antiquotation type_syntax} | + @@{ML_antiquotation const_syntax} | + @@{ML_antiquotation syntax_const}) name + "} + + \begin{description} + + \item @{command parse_translation} etc. declare syntax translation + functions to the theory. Any of these commands have a single + @{syntax text} argument that refers to an ML expression of + appropriate type, which are as follows by default: + + \medskip + {\footnotesize + \begin{tabular}{ll} + @{command parse_ast_translation} & : @{ML_type "(string * (Ast.ast list -> Ast.ast)) list"} \\ + @{command parse_translation} & : @{ML_type "(string * (term list -> term)) list"} \\ + @{command print_translation} & : @{ML_type "(string * (term list -> term)) list"} \\ + @{command typed_print_translation} & : @{ML_type "(string * (typ -> term list -> term)) list"} \\ + @{command print_ast_translation} & : @{ML_type "(string * (Ast.ast list -> Ast.ast)) list"} \\ + \end{tabular}} + \medskip + + The argument list consists of @{text "(c, tr)"} pairs, where @{text + "c"} is the syntax name of the formal entity involved, and @{text + "tr"} a function that translates a syntax form @{text "c args"} into + @{text "tr args"}. The ML naming convention for parse translations + is @{text "c_tr"} and for print translations @{text "c_tr'"}. + + The @{command_ref print_syntax} command displays the sets of names + associated with the translation functions of a theory under @{text + "parse_ast_translation"} etc. + + If the @{verbatim "("}@{keyword "advanced"}@{verbatim ")"} option is + given, the corresponding translation functions depend on the current + theory or proof context as additional argument. This allows to + implement advanced syntax mechanisms, as translations functions may + refer to specific theory declarations or auxiliary proof data. + + \item @{text "@{class_syntax c}"}, @{text "@{type_syntax c}"}, + @{text "@{const_syntax c}"} inline the authentic syntax name of the + given formal entities into the ML source. This is the + fully-qualified logical name prefixed by a special marker to + indicate its kind: thus different logical name spaces are properly + distinguished within parse trees. + + \item @{text "@{const_syntax c}"} inlines the name @{text "c"} of + the given syntax constant, having checked that it has been declared + via some @{command syntax} commands within the theory context. Note + that the usual naming convention makes syntax constants start with + underscore, to reduce the chance of accidental clashes with other + names occurring in parse trees (unqualified constants etc.). + + \end{description} +*} + + +subsubsection {* The translation strategy *} + +text {* The different kinds of translation functions are invoked during + the transformations between parse trees, ASTs and syntactic terms + (cf.\ \figref{fig:parse-print}). Whenever a combination of the form + @{text "c x\<^sub>1 \ x\<^sub>n"} is encountered, and a translation function + @{text "f"} of appropriate kind is declared for @{text "c"}, the + result is produced by evaluation of @{text "f [x\<^sub>1, \, x\<^sub>n]"} in ML. + + For AST translations, the arguments @{text "x\<^sub>1, \, x\<^sub>n"} are ASTs. A + combination has the form @{ML "Ast.Constant"}~@{text "c"} or @{ML + "Ast.Appl"}~@{text "["}@{ML Ast.Constant}~@{text "c, x\<^sub>1, \, x\<^sub>n]"}. + For term translations, the arguments are terms and a combination has + the form @{ML Const}~@{text "(c, \)"} or @{ML Const}~@{text "(c, \) + $ x\<^sub>1 $ \ $ x\<^sub>n"}. Terms allow more sophisticated transformations + than ASTs do, typically involving abstractions and bound + variables. \emph{Typed} print translations may even peek at the type + @{text "\"} of the constant they are invoked on, although that information + may be inaccurate. + + Regardless of whether they act on ASTs or terms, translation + functions called during the parsing process differ from those for + printing in their overall behaviour: + + \begin{description} + + \item [Parse translations] are applied bottom-up. The arguments are + already in translated form. The translations must not fail; + exceptions trigger an error message. There may be at most one + function associated with any syntactic name. + + \item [Print translations] are applied top-down. They are supplied + with arguments that are partly still in internal form. The result + again undergoes translation; therefore a print translation should + not introduce as head the very constant that invoked it. The + function may raise exception @{ML Match} to indicate failure; in + this event it has no effect. Multiple functions associated with + some syntactic name are tried in the order of declaration in the + theory. + + \end{description} + + Only constant atoms --- constructor @{ML Ast.Constant} for ASTs and + @{ML Const} for terms --- can invoke translation functions. This + means that parse translations can only be associated with parse tree + heads of concrete syntax, or syntactic constants introduced via + other translations. For plain identifiers within the term language, + the status of constant versus variable is not yet know during + parsing. This is in contrast to print translations, where constants + are explicitly known from the given term in its fully internal form. +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/ML_Tactic.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/ML_Tactic.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,180 @@ +theory ML_Tactic +imports Base Main +begin + +chapter {* ML tactic expressions *} + +text {* + Isar Proof methods closely resemble traditional tactics, when used + in unstructured sequences of @{command "apply"} commands. + Isabelle/Isar provides emulations for all major ML tactics of + classic Isabelle --- mostly for the sake of easy porting of existing + developments, as actual Isar proof texts would demand much less + diversity of proof methods. + + Unlike tactic expressions in ML, Isar proof methods provide proper + concrete syntax for additional arguments, options, modifiers etc. + Thus a typical method text is usually more concise than the + corresponding ML tactic. Furthermore, the Isar versions of classic + Isabelle tactics often cover several variant forms by a single + method with separate options to tune the behavior. For example, + method @{method simp} replaces all of @{ML simp_tac}~/ @{ML + asm_simp_tac}~/ @{ML full_simp_tac}~/ @{ML asm_full_simp_tac}, there + is also concrete syntax for augmenting the Simplifier context (the + current ``simpset'') in a convenient way. +*} + + +section {* Resolution tactics *} + +text {* + Classic Isabelle provides several variant forms of tactics for + single-step rule applications (based on higher-order resolution). + The space of resolution tactics has the following main dimensions. + + \begin{enumerate} + + \item The ``mode'' of resolution: intro, elim, destruct, or forward + (e.g.\ @{ML resolve_tac}, @{ML eresolve_tac}, @{ML dresolve_tac}, + @{ML forward_tac}). + + \item Optional explicit instantiation (e.g.\ @{ML resolve_tac} vs.\ + @{ML res_inst_tac}). + + \item Abbreviations for singleton arguments (e.g.\ @{ML resolve_tac} + vs.\ @{ML rtac}). + + \end{enumerate} + + Basically, the set of Isar tactic emulations @{method rule_tac}, + @{method erule_tac}, @{method drule_tac}, @{method frule_tac} (see + \secref{sec:tactics}) would be sufficient to cover the four modes, + either with or without instantiation, and either with single or + multiple arguments. Although it is more convenient in most cases to + use the plain @{method_ref (Pure) rule} method, or any of its + ``improper'' variants @{method erule}, @{method drule}, @{method + frule}. Note that explicit goal addressing is only supported by the + actual @{method rule_tac} version. + + With this in mind, plain resolution tactics correspond to Isar + methods as follows. + + \medskip + \begin{tabular}{lll} + @{ML rtac}~@{text "a 1"} & & @{text "rule a"} \\ + @{ML resolve_tac}~@{text "[a\<^sub>1, \] 1"} & & @{text "rule a\<^sub>1 \"} \\ + @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \] a 1"} & & + @{text "rule_tac x\<^sub>1 = t\<^sub>1 \ \ \ a"} \\[0.5ex] + @{ML rtac}~@{text "a i"} & & @{text "rule_tac [i] a"} \\ + @{ML resolve_tac}~@{text "[a\<^sub>1, \] i"} & & @{text "rule_tac [i] a\<^sub>1 \"} \\ + @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \] a i"} & & + @{text "rule_tac [i] x\<^sub>1 = t\<^sub>1 \ \ \ a"} \\ + \end{tabular} + \medskip + + Note that explicit goal addressing may be usually avoided by + changing the order of subgoals with @{command "defer"} or @{command + "prefer"} (see \secref{sec:tactic-commands}). +*} + + +section {* Simplifier tactics *} + +text {* + The main Simplifier tactics @{ML simp_tac} and variants (cf.\ + \cite{isabelle-ref}) are all covered by the @{method simp} and + @{method simp_all} methods (see \secref{sec:simplifier}). Note that + there is no individual goal addressing available, simplification + acts either on the first goal (@{method simp}) or all goals + (@{method simp_all}). + + \medskip + \begin{tabular}{lll} + @{ML asm_full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp} \\ + @{ML ALLGOALS}~(@{ML asm_full_simp_tac}~@{text "@{simpset}"}) & & @{method simp_all} \\[0.5ex] + @{ML simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm)"} \\ + @{ML asm_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_simp)"} \\ + @{ML full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_use)"} \\ + @{ML asm_lr_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(asm_lr)"} \\ + \end{tabular} + \medskip +*} + + +section {* Classical Reasoner tactics *} + +text {* The Classical Reasoner provides a rather large number of + variations of automated tactics, such as @{ML blast_tac}, @{ML + fast_tac}, @{ML clarify_tac} etc. The corresponding Isar methods + usually share the same base name, such as @{method blast}, @{method + fast}, @{method clarify} etc.\ (see \secref{sec:classical}). *} + + +section {* Miscellaneous tactics *} + +text {* + There are a few additional tactics defined in various theories of + Isabelle/HOL, some of these also in Isabelle/FOL or Isabelle/ZF. + The most common ones of these may be ported to Isar as follows. + + \medskip + \begin{tabular}{lll} + @{ML stac}~@{text "a 1"} & & @{text "subst a"} \\ + @{ML hyp_subst_tac}~@{text 1} & & @{text hypsubst} \\ + @{ML strip_tac}~@{text 1} & @{text "\"} & @{text "intro strip"} \\ + @{ML split_all_tac}~@{text 1} & & @{text "simp (no_asm_simp) only: split_tupled_all"} \\ + & @{text "\"} & @{text "simp only: split_tupled_all"} \\ + & @{text "\"} & @{text "clarify"} \\ + \end{tabular} +*} + + +section {* Tacticals *} + +text {* + Classic Isabelle provides a huge amount of tacticals for combination + and modification of existing tactics. This has been greatly reduced + in Isar, providing the bare minimum of combinators only: ``@{text + ","}'' (sequential composition), ``@{text "|"}'' (alternative + choices), ``@{text "?"}'' (try), ``@{text "+"}'' (repeat at least + once). These are usually sufficient in practice; if all fails, + arbitrary ML tactic code may be invoked via the @{method tactic} + method (see \secref{sec:tactics}). + + \medskip Common ML tacticals may be expressed directly in Isar as + follows: + + \medskip + \begin{tabular}{lll} + @{text "tac\<^sub>1"}~@{ML_text THEN}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1, meth\<^sub>2"} \\ + @{text "tac\<^sub>1"}~@{ML_text ORELSE}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1 | meth\<^sub>2"} \\ + @{ML TRY}~@{text tac} & & @{text "meth?"} \\ + @{ML REPEAT1}~@{text tac} & & @{text "meth+"} \\ + @{ML REPEAT}~@{text tac} & & @{text "(meth+)?"} \\ + @{ML EVERY}~@{text "[tac\<^sub>1, \]"} & & @{text "meth\<^sub>1, \"} \\ + @{ML FIRST}~@{text "[tac\<^sub>1, \]"} & & @{text "meth\<^sub>1 | \"} \\ + \end{tabular} + \medskip + + \medskip @{ML CHANGED} (see \cite{isabelle-implementation}) is + usually not required in Isar, since most basic proof methods already + fail unless there is an actual change in the goal state. + Nevertheless, ``@{text "?"}'' (try) may be used to accept + \emph{unchanged} results as well. + + \medskip @{ML ALLGOALS}, @{ML SOMEGOAL} etc.\ (see + \cite{isabelle-implementation}) are not available in Isar, since + there is no direct goal addressing. Nevertheless, some basic + methods address all goals internally, notably @{method simp_all} + (see \secref{sec:simplifier}). Also note that @{ML ALLGOALS} can be + often replaced by ``@{text "+"}'' (repeat at least once), although + this usually has a different operational behavior: subgoals are + solved in a different order. + + \medskip Iterated resolution, such as + @{ML_text "REPEAT (FIRSTGOAL (resolve_tac ...))"}, is usually better + expressed using the @{method intro} and @{method elim} methods of + Isar (see \secref{sec:classical}). +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Misc.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Misc.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,151 @@ +theory Misc +imports Base Main +begin + +chapter {* Other commands *} + +section {* Inspecting the context *} + +text {* + \begin{matharray}{rcl} + @{command_def "print_commands"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + @{command_def "print_theory"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_methods"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_attributes"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_theorems"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "find_theorems"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "find_consts"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "thm_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "unused_thms"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_facts"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_binds"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + \end{matharray} + + @{rail " + (@@{command print_theory} | @@{command print_theorems}) ('!'?) + ; + + @@{command find_theorems} ('(' @{syntax nat}? 'with_dups'? ')')? \\ (thmcriterion * ) + ; + thmcriterion: ('-'?) ('name' ':' @{syntax nameref} | 'intro' | 'elim' | 'dest' | + 'solves' | 'simp' ':' @{syntax term} | @{syntax term}) + ; + @@{command find_consts} (constcriterion * ) + ; + constcriterion: ('-'?) + ('name' ':' @{syntax nameref} | 'strict' ':' @{syntax type} | @{syntax type}) + ; + @@{command thm_deps} @{syntax thmrefs} + ; + @@{command unused_thms} ((@{syntax name} +) '-' (@{syntax name} * ))? + "} + + These commands print certain parts of the theory and proof context. + Note that there are some further ones available, such as for the set + of rules declared for simplifications. + + \begin{description} + + \item @{command "print_commands"} prints Isabelle's outer theory + syntax, including keywords and command. + + \item @{command "print_theory"} prints the main logical content of + the theory context; the ``@{text "!"}'' option indicates extra + verbosity. + + \item @{command "print_methods"} prints all proof methods + available in the current theory context. + + \item @{command "print_attributes"} prints all attributes + available in the current theory context. + + \item @{command "print_theorems"} prints theorems resulting from the + last command; the ``@{text "!"}'' option indicates extra verbosity. + + \item @{command "find_theorems"}~@{text criteria} retrieves facts + from the theory or proof context matching all of given search + criteria. The criterion @{text "name: p"} selects all theorems + whose fully qualified name matches pattern @{text p}, which may + contain ``@{text "*"}'' wildcards. The criteria @{text intro}, + @{text elim}, and @{text dest} select theorems that match the + current goal as introduction, elimination or destruction rules, + respectively. The criterion @{text "solves"} returns all rules + that would directly solve the current goal. The criterion + @{text "simp: t"} selects all rewrite rules whose left-hand side + matches the given term. The criterion term @{text t} selects all + theorems that contain the pattern @{text t} -- as usual, patterns + may contain occurrences of the dummy ``@{text _}'', schematic + variables, and type constraints. + + Criteria can be preceded by ``@{text "-"}'' to select theorems that + do \emph{not} match. Note that giving the empty list of criteria + yields \emph{all} currently known facts. An optional limit for the + number of printed facts may be given; the default is 40. By + default, duplicates are removed from the search result. Use + @{text with_dups} to display duplicates. + + \item @{command "find_consts"}~@{text criteria} prints all constants + whose type meets all of the given criteria. The criterion @{text + "strict: ty"} is met by any type that matches the type pattern + @{text ty}. Patterns may contain both the dummy type ``@{text _}'' + and sort constraints. The criterion @{text ty} is similar, but it + also matches against subtypes. The criterion @{text "name: p"} and + the prefix ``@{text "-"}'' function as described for @{command + "find_theorems"}. + + \item @{command "thm_deps"}~@{text "a\<^sub>1 \ a\<^sub>n"} + visualizes dependencies of facts, using Isabelle's graph browser + tool (see also \cite{isabelle-sys}). + + \item @{command "unused_thms"}~@{text "A\<^isub>1 \ A\<^isub>m - B\<^isub>1 \ B\<^isub>n"} + displays all unused theorems in theories @{text "B\<^isub>1 \ B\<^isub>n"} + or their parents, but not in @{text "A\<^isub>1 \ A\<^isub>m"} or their parents. + If @{text n} is @{text 0}, the end of the range of theories + defaults to the current theory. If no range is specified, + only the unused theorems in the current theory are displayed. + + \item @{command "print_facts"} prints all local facts of the + current context, both named and unnamed ones. + + \item @{command "print_binds"} prints all term abbreviations + present in the context. + + \end{description} +*} + + +section {* System commands *} + +text {* + \begin{matharray}{rcl} + @{command_def "cd"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + @{command_def "pwd"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + @{command_def "use_thy"}@{text "\<^sup>*"} & : & @{text "any \"} \\ + \end{matharray} + + @{rail " + (@@{command cd} | @@{command use_thy}) @{syntax name} + "} + + \begin{description} + + \item @{command "cd"}~@{text path} changes the current directory + of the Isabelle process. + + \item @{command "pwd"} prints the current working directory. + + \item @{command "use_thy"}~@{text A} preload theory @{text A}. + These system commands are scarcely used when working interactively, + since loading of theories is done automatically as required. + + \end{description} + + %FIXME proper place (!?) + Isabelle file specification may contain path variables (e.g.\ + @{verbatim "$ISABELLE_HOME"}) that are expanded accordingly. Note + that @{verbatim "~"} abbreviates @{verbatim "$USER_HOME"}, and + @{verbatim "~~"} abbreviates @{verbatim "$ISABELLE_HOME"}. The + general syntax for path specifications follows POSIX conventions. +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Outer_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Outer_Syntax.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,427 @@ +theory Outer_Syntax +imports Base Main +begin + +chapter {* Outer syntax --- the theory language \label{ch:outer-syntax} *} + +text {* + The rather generic framework of Isabelle/Isar syntax emerges from + three main syntactic categories: \emph{commands} of the top-level + Isar engine (covering theory and proof elements), \emph{methods} for + general goal refinements (analogous to traditional ``tactics''), and + \emph{attributes} for operations on facts (within a certain + context). Subsequently we give a reference of basic syntactic + entities underlying Isabelle/Isar syntax in a bottom-up manner. + Concrete theory and proof language elements will be introduced later + on. + + \medskip In order to get started with writing well-formed + Isabelle/Isar documents, the most important aspect to be noted is + the difference of \emph{inner} versus \emph{outer} syntax. Inner + syntax is that of Isabelle types and terms of the logic, while outer + syntax is that of Isabelle/Isar theory sources (specifications and + proofs). As a general rule, inner syntax entities may occur only as + \emph{atomic entities} within outer syntax. For example, the string + @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term + specifications within a theory, while @{verbatim "x + y"} without + quotes is not. + + Printed theory documents usually omit quotes to gain readability + (this is a matter of {\LaTeX} macro setup, say via @{verbatim + "\\isabellestyle"}, see also \cite{isabelle-sys}). Experienced + users of Isabelle/Isar may easily reconstruct the lost technical + information, while mere readers need not care about quotes at all. + + \medskip Isabelle/Isar input may contain any number of input + termination characters ``@{verbatim ";"}'' (semicolon) to separate + commands explicitly. This is particularly useful in interactive + shell sessions to make clear where the current command is intended + to end. Otherwise, the interpreter loop will continue to issue a + secondary prompt ``@{verbatim "#"}'' until an end-of-command is + clearly recognized from the input syntax, e.g.\ encounter of the + next command keyword. + + More advanced interfaces such as Proof~General \cite{proofgeneral} + do not require explicit semicolons, the amount of input text is + determined automatically by inspecting the present content of the + Emacs text buffer. In the printed presentation of Isabelle/Isar + documents semicolons are omitted altogether for readability. + + \begin{warn} + Proof~General requires certain syntax classification tables in + order to achieve properly synchronized interaction with the + Isabelle/Isar process. These tables need to be consistent with + the Isabelle version and particular logic image to be used in a + running session (common object-logics may well change the outer + syntax). The standard setup should work correctly with any of the + ``official'' logic images derived from Isabelle/HOL (including + HOLCF etc.). Users of alternative logics may need to tell + Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"} + (in conjunction with @{verbatim "-l ZF"}, to specify the default + logic image). Note that option @{verbatim "-L"} does both + of this at the same time. + \end{warn} +*} + + +section {* Lexical matters \label{sec:outer-lex} *} + +text {* The outer lexical syntax consists of three main categories of + syntax tokens: + + \begin{enumerate} + + \item \emph{major keywords} --- the command names that are available + in the present logic session; + + \item \emph{minor keywords} --- additional literal tokens required + by the syntax of commands; + + \item \emph{named tokens} --- various categories of identifiers etc. + + \end{enumerate} + + Major keywords and minor keywords are guaranteed to be disjoint. + This helps user-interfaces to determine the overall structure of a + theory text, without knowing the full details of command syntax. + Internally, there is some additional information about the kind of + major keywords, which approximates the command type (theory command, + proof command etc.). + + Keywords override named tokens. For example, the presence of a + command called @{verbatim term} inhibits the identifier @{verbatim + term}, but the string @{verbatim "\"term\""} can be used instead. + By convention, the outer syntax always allows quoted strings in + addition to identifiers, wherever a named entity is expected. + + When tokenizing a given input sequence, the lexer repeatedly takes + the longest prefix of the input that forms a valid token. Spaces, + tabs, newlines and formfeeds between tokens serve as explicit + separators. + + \medskip The categories for named tokens are defined once and for + all as follows. + + \begin{center} + \begin{supertabular}{rcl} + @{syntax_def ident} & = & @{text "letter quasiletter\<^sup>*"} \\ + @{syntax_def longident} & = & @{text "ident("}@{verbatim "."}@{text "ident)\<^sup>+"} \\ + @{syntax_def symident} & = & @{text "sym\<^sup>+ | "}@{verbatim "\\"}@{verbatim "<"}@{text ident}@{verbatim ">"} \\ + @{syntax_def nat} & = & @{text "digit\<^sup>+"} \\ + @{syntax_def float} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ + @{syntax_def var} & = & @{verbatim "?"}@{text "ident | "}@{verbatim "?"}@{text ident}@{verbatim "."}@{text nat} \\ + @{syntax_def typefree} & = & @{verbatim "'"}@{text ident} \\ + @{syntax_def typevar} & = & @{verbatim "?"}@{text "typefree | "}@{verbatim "?"}@{text typefree}@{verbatim "."}@{text nat} \\ + @{syntax_def string} & = & @{verbatim "\""} @{text "\"} @{verbatim "\""} \\ + @{syntax_def altstring} & = & @{verbatim "`"} @{text "\"} @{verbatim "`"} \\ + @{syntax_def verbatim} & = & @{verbatim "{*"} @{text "\"} @{verbatim "*"}@{verbatim "}"} \\[1ex] + + @{text letter} & = & @{text "latin | "}@{verbatim "\\"}@{verbatim "<"}@{text latin}@{verbatim ">"}@{text " | "}@{verbatim "\\"}@{verbatim "<"}@{text "latin latin"}@{verbatim ">"}@{text " | greek |"} \\ + & & @{verbatim "\<^isub>"}@{text " | "}@{verbatim "\<^isup>"} \\ + @{text quasiletter} & = & @{text "letter | digit | "}@{verbatim "_"}@{text " | "}@{verbatim "'"} \\ + @{text latin} & = & @{verbatim a}@{text " | \ | "}@{verbatim z}@{text " | "}@{verbatim A}@{text " | \ | "}@{verbatim Z} \\ + @{text digit} & = & @{verbatim "0"}@{text " | \ | "}@{verbatim "9"} \\ + @{text sym} & = & @{verbatim "!"}@{text " | "}@{verbatim "#"}@{text " | "}@{verbatim "$"}@{text " | "}@{verbatim "%"}@{text " | "}@{verbatim "&"}@{text " | "}@{verbatim "*"}@{text " | "}@{verbatim "+"}@{text " | "}@{verbatim "-"}@{text " | "}@{verbatim "/"}@{text " |"} \\ + & & @{verbatim "<"}@{text " | "}@{verbatim "="}@{text " | "}@{verbatim ">"}@{text " | "}@{verbatim "?"}@{text " | "}@{verbatim "@"}@{text " | "}@{verbatim "^"}@{text " | "}@{verbatim "_"}@{text " | "}@{verbatim "|"}@{text " | "}@{verbatim "~"} \\ + @{text greek} & = & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ + & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"} \\ + \end{supertabular} + \end{center} + + A @{syntax_ref var} or @{syntax_ref typevar} describes an unknown, + which is internally a pair of base name and index (ML type @{ML_type + indexname}). These components are either separated by a dot as in + @{text "?x.1"} or @{text "?x7.3"} or run together as in @{text + "?x1"}. The latter form is possible if the base name does not end + with digits. If the index is 0, it may be dropped altogether: + @{text "?x"} and @{text "?x0"} and @{text "?x.0"} all refer to the + same unknown, with basename @{text "x"} and index 0. + + The syntax of @{syntax_ref string} admits any characters, including + newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim + "\\"}'' (backslash) need to be escaped by a backslash; arbitrary + character codes may be specified as ``@{verbatim "\\"}@{text ddd}'', + with three decimal digits. Alternative strings according to + @{syntax_ref altstring} are analogous, using single back-quotes + instead. + + The body of @{syntax_ref verbatim} may consist of any text not + containing ``@{verbatim "*"}@{verbatim "}"}''; this allows + convenient inclusion of quotes without further escapes. There is no + way to escape ``@{verbatim "*"}@{verbatim "}"}''. If the quoted + text is {\LaTeX} source, one may usually add some blank or comment + to avoid the critical character sequence. + + Source comments take the form @{verbatim "(*"}~@{text + "\"}~@{verbatim "*)"} and may be nested, although the user-interface + might prevent this. Note that this form indicates source comments + only, which are stripped after lexical analysis of the input. The + Isar syntax also provides proper \emph{document comments} that are + considered as part of the text (see \secref{sec:comments}). + + Common mathematical symbols such as @{text \} are represented in + Isabelle as @{verbatim \}. There are infinitely many Isabelle + symbols like this, although proper presentation is left to front-end + tools such as {\LaTeX}, Proof~General, or Isabelle/jEdit. A list of + predefined Isabelle symbols that work well with these tools is given + in \appref{app:symbols}. Note that @{verbatim "\"} does not belong + to the @{text letter} category, since it is already used differently + in the Pure term language. *} + + +section {* Common syntax entities *} + +text {* + We now introduce several basic syntactic entities, such as names, + terms, and theorem specifications, which are factored out of the + actual Isar language elements to be described later. +*} + + +subsection {* Names *} + +text {* Entity @{syntax name} usually refers to any name of types, + constants, theorems etc.\ that are to be \emph{declared} or + \emph{defined} (so qualified identifiers are excluded here). Quoted + strings provide an escape for non-identifier names or those ruled + out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}). + Already existing objects are usually referenced by @{syntax + nameref}. + + @{rail " + @{syntax_def name}: @{syntax ident} | @{syntax symident} | + @{syntax string} | @{syntax nat} + ; + @{syntax_def parname}: '(' @{syntax name} ')' + ; + @{syntax_def nameref}: @{syntax name} | @{syntax longident} + "} +*} + + +subsection {* Numbers *} + +text {* The outer lexical syntax (\secref{sec:outer-lex}) admits + natural numbers and floating point numbers. These are combined as + @{syntax int} and @{syntax real} as follows. + + @{rail " + @{syntax_def int}: @{syntax nat} | '-' @{syntax nat} + ; + @{syntax_def real}: @{syntax float} | @{syntax int} + "} + + Note that there is an overlap with the category @{syntax name}, + which also includes @{syntax nat}. +*} + + +subsection {* Comments \label{sec:comments} *} + +text {* Large chunks of plain @{syntax text} are usually given + @{syntax verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim + "*"}~@{text "\"}~@{verbatim "*"}@{verbatim "}"}. For convenience, + any of the smaller text units conforming to @{syntax nameref} are + admitted as well. A marginal @{syntax comment} is of the form + @{verbatim "--"}~@{syntax text}. Any number of these may occur + within Isabelle/Isar commands. + + @{rail " + @{syntax_def text}: @{syntax verbatim} | @{syntax nameref} + ; + @{syntax_def comment}: '--' @{syntax text} + "} +*} + + +subsection {* Type classes, sorts and arities *} + +text {* + Classes are specified by plain names. Sorts have a very simple + inner syntax, which is either a single class name @{text c} or a + list @{text "{c\<^sub>1, \, c\<^sub>n}"} referring to the + intersection of these classes. The syntax of type arities is given + directly at the outer level. + + @{rail " + @{syntax_def classdecl}: @{syntax name} (('<' | '\') (@{syntax nameref} + ','))? + ; + @{syntax_def sort}: @{syntax nameref} + ; + @{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort} + "} +*} + + +subsection {* Types and terms \label{sec:types-terms} *} + +text {* + The actual inner Isabelle syntax, that of types and terms of the + logic, is far too sophisticated in order to be modelled explicitly + at the outer theory level. Basically, any such entity has to be + quoted to turn it into a single token (the parsing and type-checking + is performed internally later). For convenience, a slightly more + liberal convention is adopted: quotes may be omitted for any type or + term that is already atomic at the outer level. For example, one + may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}. + Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text + "\"} are available as well, provided these have not been superseded + by commands or other keywords already (such as @{verbatim "="} or + @{verbatim "+"}). + + @{rail " + @{syntax_def type}: @{syntax nameref} | @{syntax typefree} | + @{syntax typevar} + ; + @{syntax_def term}: @{syntax nameref} | @{syntax var} + ; + @{syntax_def prop}: @{syntax term} + "} + + Positional instantiations are indicated by giving a sequence of + terms, or the placeholder ``@{text _}'' (underscore), which means to + skip a position. + + @{rail " + @{syntax_def inst}: '_' | @{syntax term} + ; + @{syntax_def insts}: (@{syntax inst} *) + "} + + Type declarations and definitions usually refer to @{syntax + typespec} on the left-hand side. This models basic type constructor + application at the outer syntax level. Note that only plain postfix + notation is available here, but no infixes. + + @{rail " + @{syntax_def typespec}: + (() | @{syntax typefree} | '(' ( @{syntax typefree} + ',' ) ')') @{syntax name} + ; + @{syntax_def typespec_sorts}: + (() | (@{syntax typefree} ('::' @{syntax sort})?) | + '(' ( (@{syntax typefree} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name} + "} +*} + + +subsection {* Term patterns and declarations \label{sec:term-decls} *} + +text {* Wherever explicit propositions (or term fragments) occur in a + proof text, casual binding of schematic term variables may be given + specified via patterns of the form ``@{text "(\ p\<^sub>1 \ p\<^sub>n)"}''. + This works both for @{syntax term} and @{syntax prop}. + + @{rail " + @{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')' + ; + @{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')' + "} + + \medskip Declarations of local variables @{text "x :: \"} and + logical propositions @{text "a : \"} represent different views on + the same principle of introducing a local scope. In practice, one + may usually omit the typing of @{syntax vars} (due to + type-inference), and the naming of propositions (due to implicit + references of current facts). In any case, Isar proof elements + usually admit to introduce multiple such items simultaneously. + + @{rail " + @{syntax_def vars}: (@{syntax name} +) ('::' @{syntax type})? + ; + @{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +) + "} + + The treatment of multiple declarations corresponds to the + complementary focus of @{syntax vars} versus @{syntax props}. In + ``@{text "x\<^sub>1 \ x\<^sub>n :: \"}'' the typing refers to all variables, while + in @{text "a: \\<^sub>1 \ \\<^sub>n"} the naming refers to all propositions + collectively. Isar language elements that refer to @{syntax vars} + or @{syntax props} typically admit separate typings or namings via + another level of iteration, with explicit @{keyword_ref "and"} + separators; e.g.\ see @{command "fix"} and @{command "assume"} in + \secref{sec:proof-context}. +*} + + +subsection {* Attributes and theorems \label{sec:syn-att} *} + +text {* Attributes have their own ``semi-inner'' syntax, in the sense + that input conforming to @{syntax args} below is parsed by the + attribute a second time. The attribute argument specifications may + be any sequence of atomic entities (identifiers, strings etc.), or + properly bracketed argument lists. Below @{syntax atom} refers to + any atomic entity, including any @{syntax keyword} conforming to + @{syntax symident}. + + @{rail " + @{syntax_def atom}: @{syntax nameref} | @{syntax typefree} | + @{syntax typevar} | @{syntax var} | @{syntax nat} | @{syntax float} | + @{syntax keyword} + ; + arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']' + ; + @{syntax_def args}: arg * + ; + @{syntax_def attributes}: '[' (@{syntax nameref} @{syntax args} * ',') ']' + "} + + Theorem specifications come in several flavors: @{syntax axmdecl} + and @{syntax thmdecl} usually refer to axioms, assumptions or + results of goal statements, while @{syntax thmdef} collects lists of + existing theorems. Existing theorems are given by @{syntax thmref} + and @{syntax thmrefs}, the former requires an actual singleton + result. + + There are three forms of theorem references: + \begin{enumerate} + + \item named facts @{text "a"}, + + \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"}, + + \item literal fact propositions using @{syntax_ref altstring} syntax + @{verbatim "`"}@{text "\"}@{verbatim "`"} (see also method + @{method_ref fact}). + + \end{enumerate} + + Any kind of theorem specification may include lists of attributes + both on the left and right hand sides; attributes are applied to any + immediately preceding fact. If names are omitted, the theorems are + not stored within the theorem database of the theory or proof + context, but any given attributes are applied nonetheless. + + An extra pair of brackets around attributes (like ``@{text + "[[simproc a]]"}'') abbreviates a theorem reference involving an + internal dummy fact, which will be ignored later on. So only the + effect of the attribute on the background context will persist. + This form of in-place declarations is particularly useful with + commands like @{command "declare"} and @{command "using"}. + + @{rail " + @{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':' + ; + @{syntax_def thmdecl}: thmbind ':' + ; + @{syntax_def thmdef}: thmbind '=' + ; + @{syntax_def thmref}: + (@{syntax nameref} selection? | @{syntax altstring}) @{syntax attributes}? | + '[' @{syntax attributes} ']' + ; + @{syntax_def thmrefs}: @{syntax thmref} + + ; + + thmbind: @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes} + ; + selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')' + "} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Preface.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Preface.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,70 @@ +theory Preface +imports Base Main +begin + +chapter {* Preface *} + +text {* + The \emph{Isabelle} system essentially provides a generic + infrastructure for building deductive systems (programmed in + Standard ML), with a special focus on interactive theorem proving in + higher-order logics. Many years ago, even end-users would refer to + certain ML functions (goal commands, tactics, tacticals etc.) to + pursue their everyday theorem proving tasks. + + In contrast \emph{Isar} provides an interpreted language environment + of its own, which has been specifically tailored for the needs of + theory and proof development. Compared to raw ML, the Isabelle/Isar + top-level provides a more robust and comfortable development + platform, with proper support for theory development graphs, managed + transactions with unlimited undo etc. The Isabelle/Isar version of + the \emph{Proof~General} user interface + \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end + for interactive theory and proof development in this advanced + theorem proving environment, even though it is somewhat biased + towards old-style proof scripts. + + \medskip Apart from the technical advances over bare-bones ML + programming, the main purpose of the Isar language is to provide a + conceptually different view on machine-checked proofs + \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for + \emph{Intelligible semi-automated reasoning}. Drawing from both the + traditions of informal mathematical proof texts and high-level + programming languages, Isar offers a versatile environment for + structured formal proof documents. Thus properly written Isar + proofs become accessible to a broader audience than unstructured + tactic scripts (which typically only provide operational information + for the machine). Writing human-readable proof texts certainly + requires some additional efforts by the writer to achieve a good + presentation, both of formal and informal parts of the text. On the + other hand, human-readable formal texts gain some value in their own + right, independently of the mechanic proof-checking process. + + Despite its grand design of structured proof texts, Isar is able to + assimilate the old tactical style as an ``improper'' sub-language. + This provides an easy upgrade path for existing tactic scripts, as + well as some means for interactive experimentation and debugging of + structured proofs. Isabelle/Isar supports a broad range of proof + styles, both readable and unreadable ones. + + \medskip The generic Isabelle/Isar framework (see + \chref{ch:isar-framework}) works reasonably well for any Isabelle + object-logic that conforms to the natural deduction view of the + Isabelle/Pure framework. Specific language elements introduced by + Isabelle/HOL are described in \chref{ch:hol}. Although the main + language elements are already provided by the Isabelle/Pure + framework, examples given in the generic parts will usually refer to + Isabelle/HOL. + + \medskip Isar commands may be either \emph{proper} document + constructors, or \emph{improper commands}. Some proof methods and + attributes introduced later are classified as improper as well. + Improper Isar language elements, which are marked by ``@{text + "\<^sup>*"}'' in the subsequent chapters; they are often helpful + when developing proof documents, but their use is discouraged for + the final human-readable outcome. Typical examples are diagnostic + commands that print terms or theorems according to the current + context; other commands emulate old-style tactical theorem proving. +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Proof.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Proof.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1522 @@ +theory Proof +imports Base Main +begin + +chapter {* Proofs \label{ch:proofs} *} + +text {* + Proof commands perform transitions of Isar/VM machine + configurations, which are block-structured, consisting of a stack of + nodes with three main components: logical proof context, current + facts, and open goals. Isar/VM transitions are typed according to + the following three different modes of operation: + + \begin{description} + + \item @{text "proof(prove)"} means that a new goal has just been + stated that is now to be \emph{proven}; the next command may refine + it by some proof method, and enter a sub-proof to establish the + actual result. + + \item @{text "proof(state)"} is like a nested theory mode: the + context may be augmented by \emph{stating} additional assumptions, + intermediate results etc. + + \item @{text "proof(chain)"} is intermediate between @{text + "proof(state)"} and @{text "proof(prove)"}: existing facts (i.e.\ + the contents of the special ``@{fact_ref this}'' register) have been + just picked up in order to be used when refining the goal claimed + next. + + \end{description} + + The proof mode indicator may be understood as an instruction to the + writer, telling what kind of operation may be performed next. The + corresponding typings of proof commands restricts the shape of + well-formed proof texts to particular command sequences. So dynamic + arrangements of commands eventually turn out as static texts of a + certain structure. + + \Appref{ap:refcard} gives a simplified grammar of the (extensible) + language emerging that way from the different types of proof + commands. The main ideas of the overall Isar framework are + explained in \chref{ch:isar-framework}. +*} + + +section {* Proof structure *} + +subsection {* Formal notepad *} + +text {* + \begin{matharray}{rcl} + @{command_def "notepad"} & : & @{text "local_theory \ proof(state)"} \\ + \end{matharray} + + @{rail " + @@{command notepad} @'begin' + ; + @@{command end} + "} + + \begin{description} + + \item @{command "notepad"}~@{keyword "begin"} opens a proof state + without any goal statement. This allows to experiment with Isar, + without producing any persistent result. + + The notepad can be closed by @{command "end"} or discontinued by + @{command "oops"}. + + \end{description} +*} + + +subsection {* Blocks *} + +text {* + \begin{matharray}{rcl} + @{command_def "next"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "{"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "}"} & : & @{text "proof(state) \ proof(state)"} \\ + \end{matharray} + + While Isar is inherently block-structured, opening and closing + blocks is mostly handled rather casually, with little explicit + user-intervention. Any local goal statement automatically opens + \emph{two} internal blocks, which are closed again when concluding + the sub-proof (by @{command "qed"} etc.). Sections of different + context within a sub-proof may be switched via @{command "next"}, + which is just a single block-close followed by block-open again. + The effect of @{command "next"} is to reset the local proof context; + there is no goal focus involved here! + + For slightly more advanced applications, there are explicit block + parentheses as well. These typically achieve a stronger forward + style of reasoning. + + \begin{description} + + \item @{command "next"} switches to a fresh block within a + sub-proof, resetting the local context to the initial one. + + \item @{command "{"} and @{command "}"} explicitly open and close + blocks. Any current facts pass through ``@{command "{"}'' + unchanged, while ``@{command "}"}'' causes any result to be + \emph{exported} into the enclosing context. Thus fixed variables + are generalized, assumptions discharged, and local definitions + unfolded (cf.\ \secref{sec:proof-context}). There is no difference + of @{command "assume"} and @{command "presume"} in this mode of + forward reasoning --- in contrast to plain backward reasoning with + the result exported at @{command "show"} time. + + \end{description} +*} + + +subsection {* Omitting proofs *} + +text {* + \begin{matharray}{rcl} + @{command_def "oops"} & : & @{text "proof \ local_theory | theory"} \\ + \end{matharray} + + The @{command "oops"} command discontinues the current proof + attempt, while considering the partial proof text as properly + processed. This is conceptually quite different from ``faking'' + actual proofs via @{command_ref "sorry"} (see + \secref{sec:proof-steps}): @{command "oops"} does not observe the + proof structure at all, but goes back right to the theory level. + Furthermore, @{command "oops"} does not produce any result theorem + --- there is no intended claim to be able to complete the proof + in any way. + + A typical application of @{command "oops"} is to explain Isar proofs + \emph{within} the system itself, in conjunction with the document + preparation tools of Isabelle described in \chref{ch:document-prep}. + Thus partial or even wrong proof attempts can be discussed in a + logically sound manner. Note that the Isabelle {\LaTeX} macros can + be easily adapted to print something like ``@{text "\"}'' instead of + the keyword ``@{command "oops"}''. +*} + + +section {* Statements *} + +subsection {* Context elements \label{sec:proof-context} *} + +text {* + \begin{matharray}{rcl} + @{command_def "fix"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "assume"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "presume"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "def"} & : & @{text "proof(state) \ proof(state)"} \\ + \end{matharray} + + The logical proof context consists of fixed variables and + assumptions. The former closely correspond to Skolem constants, or + meta-level universal quantification as provided by the Isabelle/Pure + logical framework. Introducing some \emph{arbitrary, but fixed} + variable via ``@{command "fix"}~@{text x}'' results in a local value + that may be used in the subsequent proof as any other variable or + constant. Furthermore, any result @{text "\ \[x]"} exported from + the context will be universally closed wrt.\ @{text x} at the + outermost level: @{text "\ \x. \[x]"} (this is expressed in normal + form using Isabelle's meta-variables). + + Similarly, introducing some assumption @{text \} has two effects. + On the one hand, a local theorem is created that may be used as a + fact in subsequent proof steps. On the other hand, any result + @{text "\ \ \"} exported from the context becomes conditional wrt.\ + the assumption: @{text "\ \ \ \"}. Thus, solving an enclosing goal + using such a result would basically introduce a new subgoal stemming + from the assumption. How this situation is handled depends on the + version of assumption command used: while @{command "assume"} + insists on solving the subgoal by unification with some premise of + the goal, @{command "presume"} leaves the subgoal unchanged in order + to be proved later by the user. + + Local definitions, introduced by ``@{command "def"}~@{text "x \ + t"}'', are achieved by combining ``@{command "fix"}~@{text x}'' with + another version of assumption that causes any hypothetical equation + @{text "x \ t"} to be eliminated by the reflexivity rule. Thus, + exporting some result @{text "x \ t \ \[x]"} yields @{text "\ + \[t]"}. + + @{rail " + @@{command fix} (@{syntax vars} + @'and') + ; + (@@{command assume} | @@{command presume}) (@{syntax props} + @'and') + ; + @@{command def} (def + @'and') + ; + def: @{syntax thmdecl}? \\ @{syntax name} ('==' | '\') @{syntax term} @{syntax term_pat}? + "} + + \begin{description} + + \item @{command "fix"}~@{text x} introduces a local variable @{text + x} that is \emph{arbitrary, but fixed.} + + \item @{command "assume"}~@{text "a: \"} and @{command + "presume"}~@{text "a: \"} introduce a local fact @{text "\ \ \"} by + assumption. Subsequent results applied to an enclosing goal (e.g.\ + by @{command_ref "show"}) are handled as follows: @{command + "assume"} expects to be able to unify with existing premises in the + goal, while @{command "presume"} leaves @{text \} as new subgoals. + + Several lists of assumptions may be given (separated by + @{keyword_ref "and"}; the resulting list of current facts consists + of all of these concatenated. + + \item @{command "def"}~@{text "x \ t"} introduces a local + (non-polymorphic) definition. In results exported from the context, + @{text x} is replaced by @{text t}. Basically, ``@{command + "def"}~@{text "x \ t"}'' abbreviates ``@{command "fix"}~@{text + x}~@{command "assume"}~@{text "x \ t"}'', with the resulting + hypothetical equation solved by reflexivity. + + The default name for the definitional equation is @{text x_def}. + Several simultaneous definitions may be given at the same time. + + \end{description} + + The special name @{fact_ref prems} refers to all assumptions of the + current context as a list of theorems. This feature should be used + with great care! It is better avoided in final proof texts. +*} + + +subsection {* Term abbreviations \label{sec:term-abbrev} *} + +text {* + \begin{matharray}{rcl} + @{command_def "let"} & : & @{text "proof(state) \ proof(state)"} \\ + @{keyword_def "is"} & : & syntax \\ + \end{matharray} + + Abbreviations may be either bound by explicit @{command + "let"}~@{text "p \ t"} statements, or by annotating assumptions or + goal statements with a list of patterns ``@{text "(\ p\<^sub>1 \ + p\<^sub>n)"}''. In both cases, higher-order matching is invoked to + bind extra-logical term variables, which may be either named + schematic variables of the form @{text ?x}, or nameless dummies + ``@{variable _}'' (underscore). Note that in the @{command "let"} + form the patterns occur on the left-hand side, while the @{keyword + "is"} patterns are in postfix position. + + Polymorphism of term bindings is handled in Hindley-Milner style, + similar to ML. Type variables referring to local assumptions or + open goal statements are \emph{fixed}, while those of finished + results or bound by @{command "let"} may occur in \emph{arbitrary} + instances later. Even though actual polymorphism should be rarely + used in practice, this mechanism is essential to achieve proper + incremental type-inference, as the user proceeds to build up the + Isar proof text from left to right. + + \medskip Term abbreviations are quite different from local + definitions as introduced via @{command "def"} (see + \secref{sec:proof-context}). The latter are visible within the + logic as actual equations, while abbreviations disappear during the + input process just after type checking. Also note that @{command + "def"} does not support polymorphism. + + @{rail " + @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and') + "} + + The syntax of @{keyword "is"} patterns follows @{syntax term_pat} or + @{syntax prop_pat} (see \secref{sec:term-decls}). + + \begin{description} + + \item @{command "let"}~@{text "p\<^sub>1 = t\<^sub>1 \ \ p\<^sub>n = t\<^sub>n"} binds any + text variables in patterns @{text "p\<^sub>1, \, p\<^sub>n"} by simultaneous + higher-order matching against terms @{text "t\<^sub>1, \, t\<^sub>n"}. + + \item @{text "(\ p\<^sub>1 \ p\<^sub>n)"} resembles @{command "let"}, but + matches @{text "p\<^sub>1, \, p\<^sub>n"} against the preceding statement. Also + note that @{keyword "is"} is not a separate command, but part of + others (such as @{command "assume"}, @{command "have"} etc.). + + \end{description} + + Some \emph{implicit} term abbreviations\index{term abbreviations} + for goals and facts are available as well. For any open goal, + @{variable_ref thesis} refers to its object-level statement, + abstracted over any meta-level parameters (if present). Likewise, + @{variable_ref this} is bound for fact statements resulting from + assumptions or finished goals. In case @{variable this} refers to + an object-logic statement that is an application @{text "f t"}, then + @{text t} is bound to the special text variable ``@{variable "\"}'' + (three dots). The canonical application of this convenience are + calculational proofs (see \secref{sec:calculation}). +*} + + +subsection {* Facts and forward chaining \label{sec:proof-facts} *} + +text {* + \begin{matharray}{rcl} + @{command_def "note"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "then"} & : & @{text "proof(state) \ proof(chain)"} \\ + @{command_def "from"} & : & @{text "proof(state) \ proof(chain)"} \\ + @{command_def "with"} & : & @{text "proof(state) \ proof(chain)"} \\ + @{command_def "using"} & : & @{text "proof(prove) \ proof(prove)"} \\ + @{command_def "unfolding"} & : & @{text "proof(prove) \ proof(prove)"} \\ + \end{matharray} + + New facts are established either by assumption or proof of local + statements. Any fact will usually be involved in further proofs, + either as explicit arguments of proof methods, or when forward + chaining towards the next goal via @{command "then"} (and variants); + @{command "from"} and @{command "with"} are composite forms + involving @{command "note"}. The @{command "using"} elements + augments the collection of used facts \emph{after} a goal has been + stated. Note that the special theorem name @{fact_ref this} refers + to the most recently established facts, but only \emph{before} + issuing a follow-up claim. + + @{rail " + @@{command note} (@{syntax thmdef}? @{syntax thmrefs} + @'and') + ; + (@@{command from} | @@{command with} | @@{command using} | @@{command unfolding}) + (@{syntax thmrefs} + @'and') + "} + + \begin{description} + + \item @{command "note"}~@{text "a = b\<^sub>1 \ b\<^sub>n"} recalls existing facts + @{text "b\<^sub>1, \, b\<^sub>n"}, binding the result as @{text a}. Note that + attributes may be involved as well, both on the left and right hand + sides. + + \item @{command "then"} indicates forward chaining by the current + facts in order to establish the goal to be claimed next. The + initial proof method invoked to refine that will be offered the + facts to do ``anything appropriate'' (see also + \secref{sec:proof-steps}). For example, method @{method (Pure) rule} + (see \secref{sec:pure-meth-att}) would typically do an elimination + rather than an introduction. Automatic methods usually insert the + facts into the goal state before operation. This provides a simple + scheme to control relevance of facts in automated proof search. + + \item @{command "from"}~@{text b} abbreviates ``@{command + "note"}~@{text b}~@{command "then"}''; thus @{command "then"} is + equivalent to ``@{command "from"}~@{text this}''. + + \item @{command "with"}~@{text "b\<^sub>1 \ b\<^sub>n"} abbreviates ``@{command + "from"}~@{text "b\<^sub>1 \ b\<^sub>n \ this"}''; thus the forward chaining + is from earlier facts together with the current ones. + + \item @{command "using"}~@{text "b\<^sub>1 \ b\<^sub>n"} augments the facts being + currently indicated for use by a subsequent refinement step (such as + @{command_ref "apply"} or @{command_ref "proof"}). + + \item @{command "unfolding"}~@{text "b\<^sub>1 \ b\<^sub>n"} is structurally + similar to @{command "using"}, but unfolds definitional equations + @{text "b\<^sub>1, \ b\<^sub>n"} throughout the goal state and facts. + + \end{description} + + Forward chaining with an empty list of theorems is the same as not + chaining at all. Thus ``@{command "from"}~@{text nothing}'' has no + effect apart from entering @{text "prove(chain)"} mode, since + @{fact_ref nothing} is bound to the empty list of theorems. + + Basic proof methods (such as @{method_ref (Pure) rule}) expect multiple + facts to be given in their proper order, corresponding to a prefix + of the premises of the rule involved. Note that positions may be + easily skipped using something like @{command "from"}~@{text "_ + \ a \ b"}, for example. This involves the trivial rule + @{text "PROP \ \ PROP \"}, which is bound in Isabelle/Pure as + ``@{fact_ref "_"}'' (underscore). + + Automated methods (such as @{method simp} or @{method auto}) just + insert any given facts before their usual operation. Depending on + the kind of procedure involved, the order of facts is less + significant here. +*} + + +subsection {* Goals \label{sec:goals} *} + +text {* + \begin{matharray}{rcl} + @{command_def "lemma"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "theorem"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "corollary"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "schematic_lemma"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "schematic_theorem"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "schematic_corollary"} & : & @{text "local_theory \ proof(prove)"} \\ + @{command_def "have"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ + @{command_def "show"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ + @{command_def "hence"} & : & @{text "proof(state) \ proof(prove)"} \\ + @{command_def "thus"} & : & @{text "proof(state) \ proof(prove)"} \\ + @{command_def "print_statement"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + \end{matharray} + + From a theory context, proof mode is entered by an initial goal + command such as @{command "lemma"}, @{command "theorem"}, or + @{command "corollary"}. Within a proof, new claims may be + introduced locally as well; four variants are available here to + indicate whether forward chaining of facts should be performed + initially (via @{command_ref "then"}), and whether the final result + is meant to solve some pending goal. + + Goals may consist of multiple statements, resulting in a list of + facts eventually. A pending multi-goal is internally represented as + a meta-level conjunction (@{text "&&&"}), which is usually + split into the corresponding number of sub-goals prior to an initial + method application, via @{command_ref "proof"} + (\secref{sec:proof-steps}) or @{command_ref "apply"} + (\secref{sec:tactic-commands}). The @{method_ref induct} method + covered in \secref{sec:cases-induct} acts on multiple claims + simultaneously. + + Claims at the theory level may be either in short or long form. A + short goal merely consists of several simultaneous propositions + (often just one). A long goal includes an explicit context + specification for the subsequent conclusion, involving local + parameters and assumptions. Here the role of each part of the + statement is explicitly marked by separate keywords (see also + \secref{sec:locale}); the local assumptions being introduced here + are available as @{fact_ref assms} in the proof. Moreover, there + are two kinds of conclusions: @{element_def "shows"} states several + simultaneous propositions (essentially a big conjunction), while + @{element_def "obtains"} claims several simultaneous simultaneous + contexts of (essentially a big disjunction of eliminated parameters + and assumptions, cf.\ \secref{sec:obtain}). + + @{rail " + (@@{command lemma} | @@{command theorem} | @@{command corollary} | + @@{command schematic_lemma} | @@{command schematic_theorem} | + @@{command schematic_corollary}) @{syntax target}? (goal | longgoal) + ; + (@@{command have} | @@{command show} | @@{command hence} | @@{command thus}) goal + ; + @@{command print_statement} @{syntax modes}? @{syntax thmrefs} + ; + + goal: (@{syntax props} + @'and') + ; + longgoal: @{syntax thmdecl}? (@{syntax_ref \"includes\"}?) (@{syntax context_elem} * ) conclusion + ; + conclusion: @'shows' goal | @'obtains' (@{syntax parname}? case + '|') + ; + case: (@{syntax vars} + @'and') @'where' (@{syntax props} + @'and') + "} + + \begin{description} + + \item @{command "lemma"}~@{text "a: \"} enters proof mode with + @{text \} as main goal, eventually resulting in some fact @{text "\ + \"} to be put back into the target context. An additional @{syntax + context} specification may build up an initial proof context for the + subsequent claim; this includes local definitions and syntax as + well, see also @{syntax "includes"} in \secref{sec:bundle} and + @{syntax context_elem} in \secref{sec:locale}. + + \item @{command "theorem"}~@{text "a: \"} and @{command + "corollary"}~@{text "a: \"} are essentially the same as @{command + "lemma"}~@{text "a: \"}, but the facts are internally marked as + being of a different kind. This discrimination acts like a formal + comment. + + \item @{command "schematic_lemma"}, @{command "schematic_theorem"}, + @{command "schematic_corollary"} are similar to @{command "lemma"}, + @{command "theorem"}, @{command "corollary"}, respectively but allow + the statement to contain unbound schematic variables. + + Under normal circumstances, an Isar proof text needs to specify + claims explicitly. Schematic goals are more like goals in Prolog, + where certain results are synthesized in the course of reasoning. + With schematic statements, the inherent compositionality of Isar + proofs is lost, which also impacts performance, because proof + checking is forced into sequential mode. + + \item @{command "have"}~@{text "a: \"} claims a local goal, + eventually resulting in a fact within the current logical context. + This operation is completely independent of any pending sub-goals of + an enclosing goal statements, so @{command "have"} may be freely + used for experimental exploration of potential results within a + proof body. + + \item @{command "show"}~@{text "a: \"} is like @{command + "have"}~@{text "a: \"} plus a second stage to refine some pending + sub-goal for each one of the finished result, after having been + exported into the corresponding context (at the head of the + sub-proof of this @{command "show"} command). + + To accommodate interactive debugging, resulting rules are printed + before being applied internally. Even more, interactive execution + of @{command "show"} predicts potential failure and displays the + resulting error as a warning beforehand. Watch out for the + following message: + + %FIXME proper antiquitation + \begin{ttbox} + Problem! Local statement will fail to solve any pending goal + \end{ttbox} + + \item @{command "hence"} abbreviates ``@{command "then"}~@{command + "have"}'', i.e.\ claims a local goal to be proven by forward + chaining the current facts. Note that @{command "hence"} is also + equivalent to ``@{command "from"}~@{text this}~@{command "have"}''. + + \item @{command "thus"} abbreviates ``@{command "then"}~@{command + "show"}''. Note that @{command "thus"} is also equivalent to + ``@{command "from"}~@{text this}~@{command "show"}''. + + \item @{command "print_statement"}~@{text a} prints facts from the + current theory or proof context in long statement form, according to + the syntax for @{command "lemma"} given above. + + \end{description} + + Any goal statement causes some term abbreviations (such as + @{variable_ref "?thesis"}) to be bound automatically, see also + \secref{sec:term-abbrev}. + + The optional case names of @{element_ref "obtains"} have a twofold + meaning: (1) during the of this claim they refer to the the local + context introductions, (2) the resulting rule is annotated + accordingly to support symbolic case splits when used with the + @{method_ref cases} method (cf.\ \secref{sec:cases-induct}). +*} + + +section {* Refinement steps *} + +subsection {* Proof method expressions \label{sec:proof-meth} *} + +text {* Proof methods are either basic ones, or expressions composed + of methods via ``@{verbatim ","}'' (sequential composition), + ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}'' + (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim + "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n} + sub-goals, with default @{text "n = 1"}). In practice, proof + methods are usually just a comma separated list of @{syntax + nameref}~@{syntax args} specifications. Note that parentheses may + be dropped for single method specifications (with no arguments). + + @{rail " + @{syntax_def method}: + (@{syntax nameref} | '(' methods ')') (() | '?' | '+' | '[' @{syntax nat}? ']') + ; + methods: (@{syntax nameref} @{syntax args} | @{syntax method}) + (',' | '|') + "} + + Proper Isar proof methods do \emph{not} admit arbitrary goal + addressing, but refer either to the first sub-goal or all sub-goals + uniformly. The goal restriction operator ``@{text "[n]"}'' + evaluates a method expression within a sandbox consisting of the + first @{text n} sub-goals (which need to exist). For example, the + method ``@{text "simp_all[3]"}'' simplifies the first three + sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all + new goals that emerge from applying rule @{text "foo"} to the + originally first one. + + Improper methods, notably tactic emulations, offer a separate + low-level goal addressing scheme as explicit argument to the + individual tactic being involved. Here ``@{text "[!]"}'' refers to + all goals, and ``@{text "[n-]"}'' to all goals starting from @{text + "n"}. + + @{rail " + @{syntax_def goal_spec}: + '[' (@{syntax nat} '-' @{syntax nat} | @{syntax nat} '-' | @{syntax nat} | '!' ) ']' + "} +*} + + +subsection {* Initial and terminal proof steps \label{sec:proof-steps} *} + +text {* + \begin{matharray}{rcl} + @{command_def "proof"} & : & @{text "proof(prove) \ proof(state)"} \\ + @{command_def "qed"} & : & @{text "proof(state) \ proof(state) | local_theory | theory"} \\ + @{command_def "by"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ + @{command_def ".."} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ + @{command_def "."} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ + @{command_def "sorry"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ + \end{matharray} + + Arbitrary goal refinement via tactics is considered harmful. + Structured proof composition in Isar admits proof methods to be + invoked in two places only. + + \begin{enumerate} + + \item An \emph{initial} refinement step @{command_ref + "proof"}~@{text "m\<^sub>1"} reduces a newly stated goal to a number + of sub-goals that are to be solved later. Facts are passed to + @{text "m\<^sub>1"} for forward chaining, if so indicated by @{text + "proof(chain)"} mode. + + \item A \emph{terminal} conclusion step @{command_ref "qed"}~@{text + "m\<^sub>2"} is intended to solve remaining goals. No facts are + passed to @{text "m\<^sub>2"}. + + \end{enumerate} + + The only other (proper) way to affect pending goals in a proof body + is by @{command_ref "show"}, which involves an explicit statement of + what is to be solved eventually. Thus we avoid the fundamental + problem of unstructured tactic scripts that consist of numerous + consecutive goal transformations, with invisible effects. + + \medskip As a general rule of thumb for good proof style, initial + proof methods should either solve the goal completely, or constitute + some well-understood reduction to new sub-goals. Arbitrary + automatic proof tools that are prone leave a large number of badly + structured sub-goals are no help in continuing the proof document in + an intelligible manner. + + Unless given explicitly by the user, the default initial method is + @{method_ref (Pure) rule} (or its classical variant @{method_ref + rule}), which applies a single standard elimination or introduction + rule according to the topmost symbol involved. There is no separate + default terminal method. Any remaining goals are always solved by + assumption in the very last step. + + @{rail " + @@{command proof} method? + ; + @@{command qed} method? + ; + @@{command \"by\"} method method? + ; + (@@{command \".\"} | @@{command \"..\"} | @@{command sorry}) + "} + + \begin{description} + + \item @{command "proof"}~@{text "m\<^sub>1"} refines the goal by proof + method @{text "m\<^sub>1"}; facts for forward chaining are passed if so + indicated by @{text "proof(chain)"} mode. + + \item @{command "qed"}~@{text "m\<^sub>2"} refines any remaining goals by + proof method @{text "m\<^sub>2"} and concludes the sub-proof by assumption. + If the goal had been @{text "show"} (or @{text "thus"}), some + pending sub-goal is solved as well by the rule resulting from the + result \emph{exported} into the enclosing goal context. Thus @{text + "qed"} may fail for two reasons: either @{text "m\<^sub>2"} fails, or the + resulting rule does not fit to any pending goal\footnote{This + includes any additional ``strong'' assumptions as introduced by + @{command "assume"}.} of the enclosing context. Debugging such a + situation might involve temporarily changing @{command "show"} into + @{command "have"}, or weakening the local context by replacing + occurrences of @{command "assume"} by @{command "presume"}. + + \item @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} is a \emph{terminal + proof}\index{proof!terminal}; it abbreviates @{command + "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"}, but with + backtracking across both methods. Debugging an unsuccessful + @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} command can be done by expanding its + definition; in many cases @{command "proof"}~@{text "m\<^sub>1"} (or even + @{text "apply"}~@{text "m\<^sub>1"}) is already sufficient to see the + problem. + + \item ``@{command ".."}'' is a \emph{default + proof}\index{proof!default}; it abbreviates @{command "by"}~@{text + "rule"}. + + \item ``@{command "."}'' is a \emph{trivial + proof}\index{proof!trivial}; it abbreviates @{command "by"}~@{text + "this"}. + + \item @{command "sorry"} is a \emph{fake proof}\index{proof!fake} + pretending to solve the pending claim without further ado. This + only works in interactive development, or if the @{ML + quick_and_dirty} flag is enabled (in ML). Facts emerging from fake + proofs are not the real thing. Internally, each theorem container + is tainted by an oracle invocation, which is indicated as ``@{text + "[!]"}'' in the printed result. + + The most important application of @{command "sorry"} is to support + experimentation and top-down proof development. + + \end{description} +*} + + +subsection {* Fundamental methods and attributes \label{sec:pure-meth-att} *} + +text {* + The following proof methods and attributes refer to basic logical + operations of Isar. Further methods and attributes are provided by + several generic and object-logic specific tools and packages (see + \chref{ch:gen-tools} and \chref{ch:hol}). + + \begin{matharray}{rcl} + @{method_def "-"} & : & @{text method} \\ + @{method_def "fact"} & : & @{text method} \\ + @{method_def "assumption"} & : & @{text method} \\ + @{method_def "this"} & : & @{text method} \\ + @{method_def (Pure) "rule"} & : & @{text method} \\ + @{attribute_def (Pure) "intro"} & : & @{text attribute} \\ + @{attribute_def (Pure) "elim"} & : & @{text attribute} \\ + @{attribute_def (Pure) "dest"} & : & @{text attribute} \\ + @{attribute_def (Pure) "rule"} & : & @{text attribute} \\[0.5ex] + @{attribute_def "OF"} & : & @{text attribute} \\ + @{attribute_def "of"} & : & @{text attribute} \\ + @{attribute_def "where"} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{method fact} @{syntax thmrefs}? + ; + @@{method (Pure) rule} @{syntax thmrefs}? + ; + rulemod: ('intro' | 'elim' | 'dest') + ((('!' | () | '?') @{syntax nat}?) | 'del') ':' @{syntax thmrefs} + ; + (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) + ('!' | () | '?') @{syntax nat}? + ; + @@{attribute (Pure) rule} 'del' + ; + @@{attribute OF} @{syntax thmrefs} + ; + @@{attribute of} @{syntax insts} ('concl' ':' @{syntax insts})? + ; + @@{attribute \"where\"} + ((@{syntax name} | @{syntax var} | @{syntax typefree} | @{syntax typevar}) '=' + (@{syntax type} | @{syntax term}) * @'and') + "} + + \begin{description} + + \item ``@{method "-"}'' (minus) does nothing but insert the forward + chaining facts as premises into the goal. Note that command + @{command_ref "proof"} without any method actually performs a single + reduction step using the @{method_ref (Pure) rule} method; thus a plain + \emph{do-nothing} proof step would be ``@{command "proof"}~@{text + "-"}'' rather than @{command "proof"} alone. + + \item @{method "fact"}~@{text "a\<^sub>1 \ a\<^sub>n"} composes some fact from + @{text "a\<^sub>1, \, a\<^sub>n"} (or implicitly from the current proof context) + modulo unification of schematic type and term variables. The rule + structure is not taken into account, i.e.\ meta-level implication is + considered atomic. This is the same principle underlying literal + facts (cf.\ \secref{sec:syn-att}): ``@{command "have"}~@{text + "\"}~@{command "by"}~@{text fact}'' is equivalent to ``@{command + "note"}~@{verbatim "`"}@{text \}@{verbatim "`"}'' provided that + @{text "\ \"} is an instance of some known @{text "\ \"} in the + proof context. + + \item @{method assumption} solves some goal by a single assumption + step. All given facts are guaranteed to participate in the + refinement; this means there may be only 0 or 1 in the first place. + Recall that @{command "qed"} (\secref{sec:proof-steps}) already + concludes any remaining sub-goals by assumption, so structured + proofs usually need not quote the @{method assumption} method at + all. + + \item @{method this} applies all of the current facts directly as + rules. Recall that ``@{command "."}'' (dot) abbreviates ``@{command + "by"}~@{text this}''. + + \item @{method (Pure) rule}~@{text "a\<^sub>1 \ a\<^sub>n"} applies some rule given as + argument in backward manner; facts are used to reduce the rule + before applying it to the goal. Thus @{method (Pure) rule} without facts + is plain introduction, while with facts it becomes elimination. + + When no arguments are given, the @{method (Pure) rule} method tries to pick + appropriate rules automatically, as declared in the current context + using the @{attribute (Pure) intro}, @{attribute (Pure) elim}, + @{attribute (Pure) dest} attributes (see below). This is the + default behavior of @{command "proof"} and ``@{command ".."}'' + (double-dot) steps (see \secref{sec:proof-steps}). + + \item @{attribute (Pure) intro}, @{attribute (Pure) elim}, and + @{attribute (Pure) dest} declare introduction, elimination, and + destruct rules, to be used with method @{method (Pure) rule}, and similar + tools. Note that the latter will ignore rules declared with + ``@{text "?"}'', while ``@{text "!"}'' are used most aggressively. + + The classical reasoner (see \secref{sec:classical}) introduces its + own variants of these attributes; use qualified names to access the + present versions of Isabelle/Pure, i.e.\ @{attribute (Pure) + "Pure.intro"}. + + \item @{attribute (Pure) rule}~@{text del} undeclares introduction, + elimination, or destruct rules. + + \item @{attribute OF}~@{text "a\<^sub>1 \ a\<^sub>n"} applies some theorem to all + of the given rules @{text "a\<^sub>1, \, a\<^sub>n"} in canonical right-to-left + order, which means that premises stemming from the @{text "a\<^sub>i"} + emerge in parallel in the result, without interfering with each + other. In many practical situations, the @{text "a\<^sub>i"} do not have + premises themselves, so @{text "rule [OF a\<^sub>1 \ a\<^sub>n]"} can be actually + read as functional application (modulo unification). + + Argument positions may be effectively skipped by using ``@{text _}'' + (underscore), which refers to the propositional identity rule in the + Pure theory. + + \item @{attribute of}~@{text "t\<^sub>1 \ t\<^sub>n"} performs positional + instantiation of term variables. The terms @{text "t\<^sub>1, \, t\<^sub>n"} are + substituted for any schematic variables occurring in a theorem from + left to right; ``@{text _}'' (underscore) indicates to skip a + position. Arguments following a ``@{text "concl:"}'' specification + refer to positions of the conclusion of a rule. + + \item @{attribute "where"}~@{text "x\<^sub>1 = t\<^sub>1 \ \ x\<^sub>n = t\<^sub>n"} + performs named instantiation of schematic type and term variables + occurring in a theorem. Schematic variables have to be specified on + the left-hand side (e.g.\ @{text "?x1.3"}). The question mark may + be omitted if the variable name is a plain identifier without index. + As type instantiations are inferred from term instantiations, + explicit type instantiations are seldom necessary. + + \end{description} +*} + + +subsection {* Emulating tactic scripts \label{sec:tactic-commands} *} + +text {* + The Isar provides separate commands to accommodate tactic-style + proof scripts within the same system. While being outside the + orthodox Isar proof language, these might come in handy for + interactive exploration and debugging, or even actual tactical proof + within new-style theories (to benefit from document preparation, for + example). See also \secref{sec:tactics} for actual tactics, that + have been encapsulated as proof methods. Proper proof methods may + be used in scripts, too. + + \begin{matharray}{rcl} + @{command_def "apply"}@{text "\<^sup>*"} & : & @{text "proof(prove) \ proof(prove)"} \\ + @{command_def "apply_end"}@{text "\<^sup>*"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "done"}@{text "\<^sup>*"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ + @{command_def "defer"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ + @{command_def "prefer"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ + @{command_def "back"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ + \end{matharray} + + @{rail " + ( @@{command apply} | @@{command apply_end} ) @{syntax method} + ; + @@{command defer} @{syntax nat}? + ; + @@{command prefer} @{syntax nat} + "} + + \begin{description} + + \item @{command "apply"}~@{text m} applies proof method @{text m} in + initial position, but unlike @{command "proof"} it retains ``@{text + "proof(prove)"}'' mode. Thus consecutive method applications may be + given just as in tactic scripts. + + Facts are passed to @{text m} as indicated by the goal's + forward-chain mode, and are \emph{consumed} afterwards. Thus any + further @{command "apply"} command would always work in a purely + backward manner. + + \item @{command "apply_end"}~@{text "m"} applies proof method @{text + m} as if in terminal position. Basically, this simulates a + multi-step tactic script for @{command "qed"}, but may be given + anywhere within the proof body. + + No facts are passed to @{text m} here. Furthermore, the static + context is that of the enclosing goal (as for actual @{command + "qed"}). Thus the proof method may not refer to any assumptions + introduced in the current body, for example. + + \item @{command "done"} completes a proof script, provided that the + current goal state is solved completely. Note that actual + structured proof commands (e.g.\ ``@{command "."}'' or @{command + "sorry"}) may be used to conclude proof scripts as well. + + \item @{command "defer"}~@{text n} and @{command "prefer"}~@{text n} + shuffle the list of pending goals: @{command "defer"} puts off + sub-goal @{text n} to the end of the list (@{text "n = 1"} by + default), while @{command "prefer"} brings sub-goal @{text n} to the + front. + + \item @{command "back"} does back-tracking over the result sequence + of the latest proof command. Basically, any proof command may + return multiple results. + + \end{description} + + Any proper Isar proof method may be used with tactic script commands + such as @{command "apply"}. A few additional emulations of actual + tactics are provided as well; these would be never used in actual + structured proofs, of course. +*} + + +subsection {* Defining proof methods *} + +text {* + \begin{matharray}{rcl} + @{command_def "method_setup"} & : & @{text "theory \ theory"} \\ + \end{matharray} + + @{rail " + @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}? + ; + "} + + \begin{description} + + \item @{command "method_setup"}~@{text "name = text description"} + defines a proof method in the current theory. The given @{text + "text"} has to be an ML expression of type + @{ML_type "(Proof.context -> Proof.method) context_parser"}, cf.\ + basic parsers defined in structure @{ML_struct Args} and @{ML_struct + Attrib}. There are also combinators like @{ML METHOD} and @{ML + SIMPLE_METHOD} to turn certain tactic forms into official proof + methods; the primed versions refer to tactics with explicit goal + addressing. + + Here are some example method definitions: + + \end{description} +*} + + method_setup my_method1 = {* + Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac))) + *} "my first method (without any arguments)" + + method_setup my_method2 = {* + Scan.succeed (fn ctxt: Proof.context => + SIMPLE_METHOD' (fn i: int => no_tac)) + *} "my second method (with context)" + + method_setup my_method3 = {* + Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context => + SIMPLE_METHOD' (fn i: int => no_tac)) + *} "my third method (with theorem arguments and context)" + + +section {* Generalized elimination \label{sec:obtain} *} + +text {* + \begin{matharray}{rcl} + @{command_def "obtain"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ + @{command_def "guess"}@{text "\<^sup>*"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ + \end{matharray} + + Generalized elimination means that additional elements with certain + properties may be introduced in the current context, by virtue of a + locally proven ``soundness statement''. Technically speaking, the + @{command "obtain"} language element is like a declaration of + @{command "fix"} and @{command "assume"} (see also see + \secref{sec:proof-context}), together with a soundness proof of its + additional claim. According to the nature of existential reasoning, + assumptions get eliminated from any result exported from the context + later, provided that the corresponding parameters do \emph{not} + occur in the conclusion. + + @{rail " + @@{command obtain} @{syntax parname}? (@{syntax vars} + @'and') + @'where' (@{syntax props} + @'and') + ; + @@{command guess} (@{syntax vars} + @'and') + "} + + The derived Isar command @{command "obtain"} is defined as follows + (where @{text "b\<^sub>1, \, b\<^sub>k"} shall refer to (optional) + facts indicated for forward chaining). + \begin{matharray}{l} + @{text "\using b\<^sub>1 \ b\<^sub>k\"}~~@{command "obtain"}~@{text "x\<^sub>1 \ x\<^sub>m \ a: \\<^sub>1 \ \\<^sub>n \proof\ \"} \\[1ex] + \quad @{command "have"}~@{text "\thesis. (\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis) \ thesis"} \\ + \quad @{command "proof"}~@{method succeed} \\ + \qquad @{command "fix"}~@{text thesis} \\ + \qquad @{command "assume"}~@{text "that [Pure.intro?]: \x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis"} \\ + \qquad @{command "then"}~@{command "show"}~@{text thesis} \\ + \quad\qquad @{command "apply"}~@{text -} \\ + \quad\qquad @{command "using"}~@{text "b\<^sub>1 \ b\<^sub>k \proof\"} \\ + \quad @{command "qed"} \\ + \quad @{command "fix"}~@{text "x\<^sub>1 \ x\<^sub>m"}~@{command "assume"}@{text "\<^sup>* a: \\<^sub>1 \ \\<^sub>n"} \\ + \end{matharray} + + Typically, the soundness proof is relatively straight-forward, often + just by canonical automated tools such as ``@{command "by"}~@{text + simp}'' or ``@{command "by"}~@{text blast}''. Accordingly, the + ``@{text that}'' reduction above is declared as simplification and + introduction rule. + + In a sense, @{command "obtain"} represents at the level of Isar + proofs what would be meta-logical existential quantifiers and + conjunctions. This concept has a broad range of useful + applications, ranging from plain elimination (or introduction) of + object-level existential and conjunctions, to elimination over + results of symbolic evaluation of recursive definitions, for + example. Also note that @{command "obtain"} without parameters acts + much like @{command "have"}, where the result is treated as a + genuine assumption. + + An alternative name to be used instead of ``@{text that}'' above may + be given in parentheses. + + \medskip The improper variant @{command "guess"} is similar to + @{command "obtain"}, but derives the obtained statement from the + course of reasoning! The proof starts with a fixed goal @{text + thesis}. The subsequent proof may refine this to anything of the + form like @{text "\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ + \\<^sub>n \ thesis"}, but must not introduce new subgoals. The + final goal state is then used as reduction rule for the obtain + scheme described above. Obtained parameters @{text "x\<^sub>1, \, + x\<^sub>m"} are marked as internal by default, which prevents the + proof context from being polluted by ad-hoc variables. The variable + names and type constraints given as arguments for @{command "guess"} + specify a prefix of obtained parameters explicitly in the text. + + It is important to note that the facts introduced by @{command + "obtain"} and @{command "guess"} may not be polymorphic: any + type-variables occurring here are fixed in the present context! +*} + + +section {* Calculational reasoning \label{sec:calculation} *} + +text {* + \begin{matharray}{rcl} + @{command_def "also"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "finally"} & : & @{text "proof(state) \ proof(chain)"} \\ + @{command_def "moreover"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "ultimately"} & : & @{text "proof(state) \ proof(chain)"} \\ + @{command_def "print_trans_rules"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute trans} & : & @{text attribute} \\ + @{attribute sym} & : & @{text attribute} \\ + @{attribute symmetric} & : & @{text attribute} \\ + \end{matharray} + + Calculational proof is forward reasoning with implicit application + of transitivity rules (such those of @{text "="}, @{text "\"}, + @{text "<"}). Isabelle/Isar maintains an auxiliary fact register + @{fact_ref calculation} for accumulating results obtained by + transitivity composed with the current result. Command @{command + "also"} updates @{fact calculation} involving @{fact this}, while + @{command "finally"} exhibits the final @{fact calculation} by + forward chaining towards the next goal statement. Both commands + require valid current facts, i.e.\ may occur only after commands + that produce theorems such as @{command "assume"}, @{command + "note"}, or some finished proof of @{command "have"}, @{command + "show"} etc. The @{command "moreover"} and @{command "ultimately"} + commands are similar to @{command "also"} and @{command "finally"}, + but only collect further results in @{fact calculation} without + applying any rules yet. + + Also note that the implicit term abbreviation ``@{text "\"}'' has + its canonical application with calculational proofs. It refers to + the argument of the preceding statement. (The argument of a curried + infix expression happens to be its right-hand side.) + + Isabelle/Isar calculations are implicitly subject to block structure + in the sense that new threads of calculational reasoning are + commenced for any new block (as opened by a local goal, for + example). This means that, apart from being able to nest + calculations, there is no separate \emph{begin-calculation} command + required. + + \medskip The Isar calculation proof commands may be defined as + follows:\footnote{We suppress internal bookkeeping such as proper + handling of block-structure.} + + \begin{matharray}{rcl} + @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\ + @{command "also"}@{text "\<^sub>n+1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex] + @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] + @{command "moreover"} & \equiv & @{command "note"}~@{text "calculation = calculation this"} \\ + @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~@{text calculation} \\ + \end{matharray} + + @{rail " + (@@{command also} | @@{command finally}) ('(' @{syntax thmrefs} ')')? + ; + @@{attribute trans} (() | 'add' | 'del') + "} + + \begin{description} + + \item @{command "also"}~@{text "(a\<^sub>1 \ a\<^sub>n)"} maintains the auxiliary + @{fact calculation} register as follows. The first occurrence of + @{command "also"} in some calculational thread initializes @{fact + calculation} by @{fact this}. Any subsequent @{command "also"} on + the same level of block-structure updates @{fact calculation} by + some transitivity rule applied to @{fact calculation} and @{fact + this} (in that order). Transitivity rules are picked from the + current context, unless alternative rules are given as explicit + arguments. + + \item @{command "finally"}~@{text "(a\<^sub>1 \ a\<^sub>n)"} maintaining @{fact + calculation} in the same way as @{command "also"}, and concludes the + current calculational thread. The final result is exhibited as fact + for forward chaining towards the next goal. Basically, @{command + "finally"} just abbreviates @{command "also"}~@{command + "from"}~@{fact calculation}. Typical idioms for concluding + calculational proofs are ``@{command "finally"}~@{command + "show"}~@{text ?thesis}~@{command "."}'' and ``@{command + "finally"}~@{command "have"}~@{text \}~@{command "."}''. + + \item @{command "moreover"} and @{command "ultimately"} are + analogous to @{command "also"} and @{command "finally"}, but collect + results only, without applying rules. + + \item @{command "print_trans_rules"} prints the list of transitivity + rules (for calculational commands @{command "also"} and @{command + "finally"}) and symmetry rules (for the @{attribute symmetric} + operation and single step elimination patters) of the current + context. + + \item @{attribute trans} declares theorems as transitivity rules. + + \item @{attribute sym} declares symmetry rules, as well as + @{attribute "Pure.elim"}@{text "?"} rules. + + \item @{attribute symmetric} resolves a theorem with some rule + declared as @{attribute sym} in the current context. For example, + ``@{command "assume"}~@{text "[symmetric]: x = y"}'' produces a + swapped fact derived from that assumption. + + In structured proof texts it is often more appropriate to use an + explicit single-step elimination proof, such as ``@{command + "assume"}~@{text "x = y"}~@{command "then"}~@{command "have"}~@{text + "y = x"}~@{command ".."}''. + + \end{description} +*} + + +section {* Proof by cases and induction \label{sec:cases-induct} *} + +subsection {* Rule contexts *} + +text {* + \begin{matharray}{rcl} + @{command_def "case"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "print_cases"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def case_names} & : & @{text attribute} \\ + @{attribute_def case_conclusion} & : & @{text attribute} \\ + @{attribute_def params} & : & @{text attribute} \\ + @{attribute_def consumes} & : & @{text attribute} \\ + \end{matharray} + + The puristic way to build up Isar proof contexts is by explicit + language elements like @{command "fix"}, @{command "assume"}, + @{command "let"} (see \secref{sec:proof-context}). This is adequate + for plain natural deduction, but easily becomes unwieldy in concrete + verification tasks, which typically involve big induction rules with + several cases. + + The @{command "case"} command provides a shorthand to refer to a + local context symbolically: certain proof methods provide an + environment of named ``cases'' of the form @{text "c: x\<^sub>1, \, + x\<^sub>m, \\<^sub>1, \, \\<^sub>n"}; the effect of ``@{command + "case"}~@{text c}'' is then equivalent to ``@{command "fix"}~@{text + "x\<^sub>1 \ x\<^sub>m"}~@{command "assume"}~@{text "c: \\<^sub>1 \ + \\<^sub>n"}''. Term bindings may be covered as well, notably + @{variable ?case} for the main conclusion. + + By default, the ``terminology'' @{text "x\<^sub>1, \, x\<^sub>m"} of + a case value is marked as hidden, i.e.\ there is no way to refer to + such parameters in the subsequent proof text. After all, original + rule parameters stem from somewhere outside of the current proof + text. By using the explicit form ``@{command "case"}~@{text "(c + y\<^sub>1 \ y\<^sub>m)"}'' instead, the proof author is able to + chose local names that fit nicely into the current context. + + \medskip It is important to note that proper use of @{command + "case"} does not provide means to peek at the current goal state, + which is not directly observable in Isar! Nonetheless, goal + refinement commands do provide named cases @{text "goal\<^sub>i"} + for each subgoal @{text "i = 1, \, n"} of the resulting goal state. + Using this extra feature requires great care, because some bits of + the internal tactical machinery intrude the proof text. In + particular, parameter names stemming from the left-over of automated + reasoning tools are usually quite unpredictable. + + Under normal circumstances, the text of cases emerge from standard + elimination or induction rules, which in turn are derived from + previous theory specifications in a canonical way (say from + @{command "inductive"} definitions). + + \medskip Proper cases are only available if both the proof method + and the rules involved support this. By using appropriate + attributes, case names, conclusions, and parameters may be also + declared by hand. Thus variant versions of rules that have been + derived manually become ready to use in advanced case analysis + later. + + @{rail " + @@{command case} (caseref | '(' caseref (('_' | @{syntax name}) +) ')') + ; + caseref: nameref attributes? + ; + + @@{attribute case_names} ((@{syntax name} ( '[' (('_' | @{syntax name}) +) ']' ) ? ) +) + ; + @@{attribute case_conclusion} @{syntax name} (@{syntax name} * ) + ; + @@{attribute params} ((@{syntax name} * ) + @'and') + ; + @@{attribute consumes} @{syntax nat}? + "} + + \begin{description} + + \item @{command "case"}~@{text "(c x\<^sub>1 \ x\<^sub>m)"} invokes a named local + context @{text "c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>m"}, as provided by an + appropriate proof method (such as @{method_ref cases} and + @{method_ref induct}). The command ``@{command "case"}~@{text "(c + x\<^sub>1 \ x\<^sub>m)"}'' abbreviates ``@{command "fix"}~@{text "x\<^sub>1 \ + x\<^sub>m"}~@{command "assume"}~@{text "c: \\<^sub>1 \ \\<^sub>n"}''. + + \item @{command "print_cases"} prints all local contexts of the + current state, using Isar proof language notation. + + \item @{attribute case_names}~@{text "c\<^sub>1 \ c\<^sub>k"} declares names for + the local contexts of premises of a theorem; @{text "c\<^sub>1, \, c\<^sub>k"} + refers to the \emph{prefix} of the list of premises. Each of the + cases @{text "c\<^isub>i"} can be of the form @{text "c[h\<^isub>1 \ h\<^isub>n]"} where + the @{text "h\<^isub>1 \ h\<^isub>n"} are the names of the hypotheses in case @{text "c\<^isub>i"} + from left to right. + + \item @{attribute case_conclusion}~@{text "c d\<^sub>1 \ d\<^sub>k"} declares + names for the conclusions of a named premise @{text c}; here @{text + "d\<^sub>1, \, d\<^sub>k"} refers to the prefix of arguments of a logical formula + built by nesting a binary connective (e.g.\ @{text "\"}). + + Note that proof methods such as @{method induct} and @{method + coinduct} already provide a default name for the conclusion as a + whole. The need to name subformulas only arises with cases that + split into several sub-cases, as in common co-induction rules. + + \item @{attribute params}~@{text "p\<^sub>1 \ p\<^sub>m \ \ q\<^sub>1 \ q\<^sub>n"} renames + the innermost parameters of premises @{text "1, \, n"} of some + theorem. An empty list of names may be given to skip positions, + leaving the present parameters unchanged. + + Note that the default usage of case rules does \emph{not} directly + expose parameters to the proof context. + + \item @{attribute consumes}~@{text n} declares the number of ``major + premises'' of a rule, i.e.\ the number of facts to be consumed when + it is applied by an appropriate proof method. The default value of + @{attribute consumes} is @{text "n = 1"}, which is appropriate for + the usual kind of cases and induction rules for inductive sets (cf.\ + \secref{sec:hol-inductive}). Rules without any @{attribute + consumes} declaration given are treated as if @{attribute + consumes}~@{text 0} had been specified. + + Note that explicit @{attribute consumes} declarations are only + rarely needed; this is already taken care of automatically by the + higher-level @{attribute cases}, @{attribute induct}, and + @{attribute coinduct} declarations. + + \end{description} +*} + + +subsection {* Proof methods *} + +text {* + \begin{matharray}{rcl} + @{method_def cases} & : & @{text method} \\ + @{method_def induct} & : & @{text method} \\ + @{method_def induction} & : & @{text method} \\ + @{method_def coinduct} & : & @{text method} \\ + \end{matharray} + + The @{method cases}, @{method induct}, @{method induction}, + and @{method coinduct} + methods provide a uniform interface to common proof techniques over + datatypes, inductive predicates (or sets), recursive functions etc. + The corresponding rules may be specified and instantiated in a + casual manner. Furthermore, these methods provide named local + contexts that may be invoked via the @{command "case"} proof command + within the subsequent proof text. This accommodates compact proof + texts even when reasoning about large specifications. + + The @{method induct} method also provides some additional + infrastructure in order to be applicable to structure statements + (either using explicit meta-level connectives, or including facts + and parameters separately). This avoids cumbersome encoding of + ``strengthened'' inductive statements within the object-logic. + + Method @{method induction} differs from @{method induct} only in + the names of the facts in the local context invoked by the @{command "case"} + command. + + @{rail " + @@{method cases} ('(' 'no_simp' ')')? \\ + (@{syntax insts} * @'and') rule? + ; + (@@{method induct} | @@{method induction}) ('(' 'no_simp' ')')? (definsts * @'and') \\ arbitrary? taking? rule? + ; + @@{method coinduct} @{syntax insts} taking rule? + ; + + rule: ('type' | 'pred' | 'set') ':' (@{syntax nameref} +) | 'rule' ':' (@{syntax thmref} +) + ; + definst: @{syntax name} ('==' | '\') @{syntax term} | '(' @{syntax term} ')' | @{syntax inst} + ; + definsts: ( definst * ) + ; + arbitrary: 'arbitrary' ':' ((@{syntax term} * ) @'and' +) + ; + taking: 'taking' ':' @{syntax insts} + "} + + \begin{description} + + \item @{method cases}~@{text "insts R"} applies method @{method + rule} with an appropriate case distinction theorem, instantiated to + the subjects @{text insts}. Symbolic case names are bound according + to the rule's local contexts. + + The rule is determined as follows, according to the facts and + arguments passed to the @{method cases} method: + + \medskip + \begin{tabular}{llll} + facts & & arguments & rule \\\hline + & @{method cases} & & classical case split \\ + & @{method cases} & @{text t} & datatype exhaustion (type of @{text t}) \\ + @{text "\ A t"} & @{method cases} & @{text "\"} & inductive predicate/set elimination (of @{text A}) \\ + @{text "\"} & @{method cases} & @{text "\ rule: R"} & explicit rule @{text R} \\ + \end{tabular} + \medskip + + Several instantiations may be given, referring to the \emph{suffix} + of premises of the case rule; within each premise, the \emph{prefix} + of variables is instantiated. In most situations, only a single + term needs to be specified; this refers to the first variable of the + last premise (it is usually the same for all cases). The @{text + "(no_simp)"} option can be used to disable pre-simplification of + cases (see the description of @{method induct} below for details). + + \item @{method induct}~@{text "insts R"} and + @{method induction}~@{text "insts R"} are analogous to the + @{method cases} method, but refer to induction rules, which are + determined as follows: + + \medskip + \begin{tabular}{llll} + facts & & arguments & rule \\\hline + & @{method induct} & @{text "P x"} & datatype induction (type of @{text x}) \\ + @{text "\ A x"} & @{method induct} & @{text "\"} & predicate/set induction (of @{text A}) \\ + @{text "\"} & @{method induct} & @{text "\ rule: R"} & explicit rule @{text R} \\ + \end{tabular} + \medskip + + Several instantiations may be given, each referring to some part of + a mutual inductive definition or datatype --- only related partial + induction rules may be used together, though. Any of the lists of + terms @{text "P, x, \"} refers to the \emph{suffix} of variables + present in the induction rule. This enables the writer to specify + only induction variables, or both predicates and variables, for + example. + + Instantiations may be definitional: equations @{text "x \ t"} + introduce local definitions, which are inserted into the claim and + discharged after applying the induction rule. Equalities reappear + in the inductive cases, but have been transformed according to the + induction principle being involved here. In order to achieve + practically useful induction hypotheses, some variables occurring in + @{text t} need to be fixed (see below). Instantiations of the form + @{text t}, where @{text t} is not a variable, are taken as a + shorthand for \mbox{@{text "x \ t"}}, where @{text x} is a fresh + variable. If this is not intended, @{text t} has to be enclosed in + parentheses. By default, the equalities generated by definitional + instantiations are pre-simplified using a specific set of rules, + usually consisting of distinctness and injectivity theorems for + datatypes. This pre-simplification may cause some of the parameters + of an inductive case to disappear, or may even completely delete + some of the inductive cases, if one of the equalities occurring in + their premises can be simplified to @{text False}. The @{text + "(no_simp)"} option can be used to disable pre-simplification. + Additional rules to be used in pre-simplification can be declared + using the @{attribute_def induct_simp} attribute. + + The optional ``@{text "arbitrary: x\<^sub>1 \ x\<^sub>m"}'' + specification generalizes variables @{text "x\<^sub>1, \, + x\<^sub>m"} of the original goal before applying induction. One can + separate variables by ``@{text "and"}'' to generalize them in other + goals then the first. Thus induction hypotheses may become + sufficiently general to get the proof through. Together with + definitional instantiations, one may effectively perform induction + over expressions of a certain structure. + + The optional ``@{text "taking: t\<^sub>1 \ t\<^sub>n"}'' + specification provides additional instantiations of a prefix of + pending variables in the rule. Such schematic induction rules + rarely occur in practice, though. + + \item @{method coinduct}~@{text "inst R"} is analogous to the + @{method induct} method, but refers to coinduction rules, which are + determined as follows: + + \medskip + \begin{tabular}{llll} + goal & & arguments & rule \\\hline + & @{method coinduct} & @{text x} & type coinduction (type of @{text x}) \\ + @{text "A x"} & @{method coinduct} & @{text "\"} & predicate/set coinduction (of @{text A}) \\ + @{text "\"} & @{method coinduct} & @{text "\ rule: R"} & explicit rule @{text R} \\ + \end{tabular} + + Coinduction is the dual of induction. Induction essentially + eliminates @{text "A x"} towards a generic result @{text "P x"}, + while coinduction introduces @{text "A x"} starting with @{text "B + x"}, for a suitable ``bisimulation'' @{text B}. The cases of a + coinduct rule are typically named after the predicates or sets being + covered, while the conclusions consist of several alternatives being + named after the individual destructor patterns. + + The given instantiation refers to the \emph{suffix} of variables + occurring in the rule's major premise, or conclusion if unavailable. + An additional ``@{text "taking: t\<^sub>1 \ t\<^sub>n"}'' + specification may be required in order to specify the bisimulation + to be used in the coinduction step. + + \end{description} + + Above methods produce named local contexts, as determined by the + instantiated rule as given in the text. Beyond that, the @{method + induct} and @{method coinduct} methods guess further instantiations + from the goal specification itself. Any persisting unresolved + schematic variables of the resulting rule will render the the + corresponding case invalid. The term binding @{variable ?case} for + the conclusion will be provided with each case, provided that term + is fully specified. + + The @{command "print_cases"} command prints all named cases present + in the current proof state. + + \medskip Despite the additional infrastructure, both @{method cases} + and @{method coinduct} merely apply a certain rule, after + instantiation, while conforming due to the usual way of monotonic + natural deduction: the context of a structured statement @{text + "\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ \"} + reappears unchanged after the case split. + + The @{method induct} method is fundamentally different in this + respect: the meta-level structure is passed through the + ``recursive'' course involved in the induction. Thus the original + statement is basically replaced by separate copies, corresponding to + the induction hypotheses and conclusion; the original goal context + is no longer available. Thus local assumptions, fixed parameters + and definitions effectively participate in the inductive rephrasing + of the original statement. + + In @{method induct} proofs, local assumptions introduced by cases are split + into two different kinds: @{text hyps} stemming from the rule and + @{text prems} from the goal statement. This is reflected in the + extracted cases accordingly, so invoking ``@{command "case"}~@{text + c}'' will provide separate facts @{text c.hyps} and @{text c.prems}, + as well as fact @{text c} to hold the all-inclusive list. + + In @{method induction} proofs, local assumptions introduced by cases are + split into three different kinds: @{text IH}, the induction hypotheses, + @{text hyps}, the remaining hypotheses stemming from the rule, and + @{text prems}, the assumptions from the goal statement. The names are + @{text c.IH}, @{text c.hyps} and @{text c.prems}, as above. + + + \medskip Facts presented to either method are consumed according to + the number of ``major premises'' of the rule involved, which is + usually 0 for plain cases and induction rules of datatypes etc.\ and + 1 for rules of inductive predicates or sets and the like. The + remaining facts are inserted into the goal verbatim before the + actual @{text cases}, @{text induct}, or @{text coinduct} rule is + applied. +*} + + +subsection {* Declaring rules *} + +text {* + \begin{matharray}{rcl} + @{command_def "print_induct_rules"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{attribute_def cases} & : & @{text attribute} \\ + @{attribute_def induct} & : & @{text attribute} \\ + @{attribute_def coinduct} & : & @{text attribute} \\ + \end{matharray} + + @{rail " + @@{attribute cases} spec + ; + @@{attribute induct} spec + ; + @@{attribute coinduct} spec + ; + + spec: (('type' | 'pred' | 'set') ':' @{syntax nameref}) | 'del' + "} + + \begin{description} + + \item @{command "print_induct_rules"} prints cases and induct rules + for predicates (or sets) and types of the current context. + + \item @{attribute cases}, @{attribute induct}, and @{attribute + coinduct} (as attributes) declare rules for reasoning about + (co)inductive predicates (or sets) and types, using the + corresponding methods of the same name. Certain definitional + packages of object-logics usually declare emerging cases and + induction rules as expected, so users rarely need to intervene. + + Rules may be deleted via the @{text "del"} specification, which + covers all of the @{text "type"}/@{text "pred"}/@{text "set"} + sub-categories simultaneously. For example, @{attribute + cases}~@{text del} removes any @{attribute cases} rules declared for + some type, predicate, or set. + + Manual rule declarations usually refer to the @{attribute + case_names} and @{attribute params} attributes to adjust names of + cases and parameters of a rule; the @{attribute consumes} + declaration is taken care of automatically: @{attribute + consumes}~@{text 0} is specified for ``type'' rules and @{attribute + consumes}~@{text 1} for ``predicate'' / ``set'' rules. + + \end{description} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Quick_Reference.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Quick_Reference.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,231 @@ +theory Quick_Reference +imports Base Main +begin + +chapter {* Isabelle/Isar quick reference \label{ap:refcard} *} + +section {* Proof commands *} + +subsection {* Primitives and basic syntax *} + +text {* + \begin{tabular}{ll} + @{command "fix"}~@{text x} & augment context by @{text "\x. \"} \\ + @{command "assume"}~@{text "a: \"} & augment context by @{text "\ \ \"} \\ + @{command "then"} & indicate forward chaining of facts \\ + @{command "have"}~@{text "a: \"} & prove local result \\ + @{command "show"}~@{text "a: \"} & prove local result, refining some goal \\ + @{command "using"}~@{text a} & indicate use of additional facts \\ + @{command "unfolding"}~@{text a} & unfold definitional equations \\ + @{command "proof"}~@{text "m\<^sub>1"}~\dots~@{command "qed"}~@{text "m\<^sub>2"} & indicate proof structure and refinements \\ + @{command "{"}~@{text "\"}~@{command "}"} & indicate explicit blocks \\ + @{command "next"} & switch blocks \\ + @{command "note"}~@{text "a = b"} & reconsider facts \\ + @{command "let"}~@{text "p = t"} & abbreviate terms by higher-order matching \\ + @{command "write"}~@{text "c (mx)"} & declare local mixfix syntax \\ + \end{tabular} + + \medskip + + \begin{tabular}{rcl} + @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\ + & @{text "|"} & @{text "prfx\<^sup>*"}~@{command "done"} \\ + @{text prfx} & = & @{command "apply"}~@{text method} \\ + & @{text "|"} & @{command "using"}~@{text "facts"} \\ + & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ + @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ + & @{text "|"} & @{command "next"} \\ + & @{text "|"} & @{command "note"}~@{text "name = facts"} \\ + & @{text "|"} & @{command "let"}~@{text "term = term"} \\ + & @{text "|"} & @{command "write"}~@{text "name (mixfix)"} \\ + & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ + & @{text "|"} & @{command "assume"}~@{text "name: props"} \\ + & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ + @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ + & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ + \end{tabular} +*} + + +subsection {* Abbreviations and synonyms *} + +text {* + \begin{tabular}{rcl} + @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} & @{text "\"} & + @{command "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"} \\ + @{command ".."} & @{text "\"} & @{command "by"}~@{text rule} \\ + @{command "."} & @{text "\"} & @{command "by"}~@{text this} \\ + @{command "hence"} & @{text "\"} & @{command "then"}~@{command "have"} \\ + @{command "thus"} & @{text "\"} & @{command "then"}~@{command "show"} \\ + @{command "from"}~@{text a} & @{text "\"} & @{command "note"}~@{text a}~@{command "then"} \\ + @{command "with"}~@{text a} & @{text "\"} & @{command "from"}~@{text "a \ this"} \\ + @{command "from"}~@{text this} & @{text "\"} & @{command "then"} \\ + @{command "from"}~@{text this}~@{command "have"} & @{text "\"} & @{command "hence"} \\ + @{command "from"}~@{text this}~@{command "show"} & @{text "\"} & @{command "thus"} \\ + \end{tabular} +*} + + +subsection {* Derived elements *} + +text {* + \begin{tabular}{rcl} + @{command "also"}@{text "\<^sub>0"} & @{text "\"} & + @{command "note"}~@{text "calculation = this"} \\ + @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & @{text "\"} & + @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\ + @{command "finally"} & @{text "\"} & + @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] + @{command "moreover"} & @{text "\"} & + @{command "note"}~@{text "calculation = calculation this"} \\ + @{command "ultimately"} & @{text "\"} & + @{command "moreover"}~@{command "from"}~@{text calculation} \\[0.5ex] + @{command "presume"}~@{text "a: \"} & @{text "\"} & + @{command "assume"}~@{text "a: \"} \\ + @{command "def"}~@{text "a: x \ t"} & @{text "\"} & + @{command "fix"}~@{text x}~@{command "assume"}~@{text "a: x \ t"} \\ + @{command "obtain"}~@{text "x \ a: \"} & @{text "\"} & + @{text "\"}~@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: \"} \\ + @{command "case"}~@{text c} & @{text "\"} & + @{command "fix"}~@{text x}~@{command "assume"}~@{text "c: \"} \\ + @{command "sorry"} & @{text "\"} & + @{command "by"}~@{text cheating} \\ + \end{tabular} +*} + + +subsection {* Diagnostic commands *} + +text {* + \begin{tabular}{ll} + @{command "pr"} & print current state \\ + @{command "thm"}~@{text a} & print fact \\ + @{command "prop"}~@{text \} & print proposition \\ + @{command "term"}~@{text t} & print term \\ + @{command "typ"}~@{text \} & print type \\ + \end{tabular} +*} + + +section {* Proof methods *} + +text {* + \begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] + @{method assumption} & apply some assumption \\ + @{method this} & apply current facts \\ + @{method rule}~@{text a} & apply some rule \\ + @{method rule} & apply standard rule (default for @{command "proof"}) \\ + @{method contradiction} & apply @{text "\"} elimination rule (any order) \\ + @{method cases}~@{text t} & case analysis (provides cases) \\ + @{method induct}~@{text x} & proof by induction (provides cases) \\[2ex] + + \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] + @{method "-"} & no rules \\ + @{method intro}~@{text a} & introduction rules \\ + @{method intro_classes} & class introduction rules \\ + @{method elim}~@{text a} & elimination rules \\ + @{method unfold}~@{text a} & definitional rewrite rules \\[2ex] + + \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] + @{method iprover} & intuitionistic proof search \\ + @{method blast}, @{method fast} & Classical Reasoner \\ + @{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\ + @{method auto}, @{method force} & Simplifier + Classical Reasoner \\ + @{method arith} & Arithmetic procedures \\ + \end{tabular} +*} + + +section {* Attributes *} + +text {* + \begin{tabular}{ll} + \multicolumn{2}{l}{\textbf{Rules}} \\[0.5ex] + @{attribute OF}~@{text a} & rule resolved with facts (skipping ``@{text _}'') \\ + @{attribute of}~@{text t} & rule instantiated with terms (skipping ``@{text _}'') \\ + @{attribute "where"}~@{text "x = t"} & rule instantiated with terms, by variable name \\ + @{attribute symmetric} & resolution with symmetry rule \\ + @{attribute THEN}~@{text b} & resolution with another rule \\ + @{attribute rule_format} & result put into standard rule format \\ + @{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex] + + \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] + @{attribute simp} & Simplifier rule \\ + @{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\ + @{attribute iff} & Simplifier + Classical Reasoner rule \\ + @{attribute split} & case split rule \\ + @{attribute trans} & transitivity rule \\ + @{attribute sym} & symmetry rule \\ + \end{tabular} +*} + + +section {* Rule declarations and methods *} + +text {* + \begin{tabular}{l|lllll} + & @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\ + & & & @{method fast} & @{method simp_all} & @{method force} \\ + \hline + @{attribute Pure.elim}@{text "!"} @{attribute Pure.intro}@{text "!"} + & @{text "\"} & @{text "\"} \\ + @{attribute Pure.elim} @{attribute Pure.intro} + & @{text "\"} & @{text "\"} \\ + @{attribute elim}@{text "!"} @{attribute intro}@{text "!"} + & @{text "\"} & & @{text "\"} & & @{text "\"} \\ + @{attribute elim} @{attribute intro} + & @{text "\"} & & @{text "\"} & & @{text "\"} \\ + @{attribute iff} + & @{text "\"} & & @{text "\"} & @{text "\"} & @{text "\"} \\ + @{attribute iff}@{text "?"} + & @{text "\"} \\ + @{attribute elim}@{text "?"} @{attribute intro}@{text "?"} + & @{text "\"} \\ + @{attribute simp} + & & & & @{text "\"} & @{text "\"} \\ + @{attribute cong} + & & & & @{text "\"} & @{text "\"} \\ + @{attribute split} + & & & & @{text "\"} & @{text "\"} \\ + \end{tabular} +*} + + +section {* Emulating tactic scripts *} + +subsection {* Commands *} + +text {* + \begin{tabular}{ll} + @{command "apply"}~@{text m} & apply proof method at initial position \\ + @{command "apply_end"}~@{text m} & apply proof method near terminal position \\ + @{command "done"} & complete proof \\ + @{command "defer"}~@{text n} & move subgoal to end \\ + @{command "prefer"}~@{text n} & move subgoal to beginning \\ + @{command "back"} & backtrack last command \\ + \end{tabular} +*} + + +subsection {* Methods *} + +text {* + \begin{tabular}{ll} + @{method rule_tac}~@{text insts} & resolution (with instantiation) \\ + @{method erule_tac}~@{text insts} & elim-resolution (with instantiation) \\ + @{method drule_tac}~@{text insts} & destruct-resolution (with instantiation) \\ + @{method frule_tac}~@{text insts} & forward-resolution (with instantiation) \\ + @{method cut_tac}~@{text insts} & insert facts (with instantiation) \\ + @{method thin_tac}~@{text \} & delete assumptions \\ + @{method subgoal_tac}~@{text \} & new claims \\ + @{method rename_tac}~@{text x} & rename innermost goal parameters \\ + @{method rotate_tac}~@{text n} & rotate assumptions of goal \\ + @{method tactic}~@{text "text"} & arbitrary ML tactic \\ + @{method case_tac}~@{text t} & exhaustion (datatypes) \\ + @{method induct_tac}~@{text x} & induction (datatypes) \\ + @{method ind_cases}~@{text t} & exhaustion + simplification (inductive predicates) \\ + \end{tabular} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Spec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Spec.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1337 @@ +theory Spec +imports Base Main +begin + +chapter {* Specifications *} + +text {* + The Isabelle/Isar theory format integrates specifications and + proofs, supporting interactive development with unlimited undo + operation. There is an integrated document preparation system (see + \chref{ch:document-prep}), for typesetting formal developments + together with informal text. The resulting hyper-linked PDF + documents can be used both for WWW presentation and printed copies. + + The Isar proof language (see \chref{ch:proofs}) is embedded into the + theory language as a proper sub-language. Proof mode is entered by + stating some @{command theorem} or @{command lemma} at the theory + level, and left again with the final conclusion (e.g.\ via @{command + qed}). Some theory specification mechanisms also require a proof, + such as @{command typedef} in HOL, which demands non-emptiness of + the representing sets. +*} + + +section {* Defining theories \label{sec:begin-thy} *} + +text {* + \begin{matharray}{rcl} + @{command_def "theory"} & : & @{text "toplevel \ theory"} \\ + @{command_def (global) "end"} & : & @{text "theory \ toplevel"} \\ + \end{matharray} + + Isabelle/Isar theories are defined via theory files, which may + contain both specifications and proofs; occasionally definitional + mechanisms also require some explicit proof. The theory body may be + sub-structured by means of \emph{local theory targets}, such as + @{command "locale"} and @{command "class"}. + + The first proper command of a theory is @{command "theory"}, which + indicates imports of previous theories and optional dependencies on + other source files (usually in ML). Just preceding the initial + @{command "theory"} command there may be an optional @{command + "header"} declaration, which is only relevant to document + preparation: see also the other section markup commands in + \secref{sec:markup}. + + A theory is concluded by a final @{command (global) "end"} command, + one that does not belong to a local theory target. No further + commands may follow such a global @{command (global) "end"}, + although some user-interfaces might pretend that trailing input is + admissible. + + @{rail " + @@{command theory} @{syntax name} imports \\ keywords? uses? @'begin' + ; + imports: @'imports' (@{syntax name} +) + ; + keywords: @'keywords' ((@{syntax string} +) ('::' @{syntax name} @{syntax tags})? + @'and') + ; + uses: @'uses' ((@{syntax name} | @{syntax parname}) +) + "} + + \begin{description} + + \item @{command "theory"}~@{text "A \ B\<^sub>1 \ B\<^sub>n \"} + starts a new theory @{text A} based on the merge of existing + theories @{text "B\<^sub>1 \ B\<^sub>n"}. Due to the possibility to import more + than one ancestor, the resulting theory structure of an Isabelle + session forms a directed acyclic graph (DAG). Isabelle takes care + that sources contributing to the development graph are always + up-to-date: changed files are automatically rechecked whenever a + theory header specification is processed. + + The optional @{keyword_def "keywords"} specification declares outer + syntax (\chref{ch:outer-syntax}) that is introduced in this theory + later on (rare in end-user applications). Both minor keywords and + major keywords of the Isar command language need to be specified, in + order to make parsing of proof documents work properly. Command + keywords need to be classified according to their structural role in + the formal text. Examples may be seen in Isabelle/HOL sources + itself, such as @{keyword "keywords"}~@{verbatim "\"typedef\""} + @{text ":: thy_goal"} or @{keyword "keywords"}~@{verbatim + "\"datatype\""} @{text ":: thy_decl"} for theory-level declarations + with and without proof, respectively. Additional @{syntax tags} + provide defaults for document preparation (\secref{sec:tags}). + + The optional @{keyword_def "uses"} specification declares additional + dependencies on external files (notably ML sources). Files will be + loaded immediately (as ML), unless the name is parenthesized. The + latter case records a dependency that needs to be resolved later in + the text, usually via explicit @{command_ref "use"} for ML files; + other file formats require specific load commands defined by the + corresponding tools or packages. + + \item @{command (global) "end"} concludes the current theory + definition. Note that some other commands, e.g.\ local theory + targets @{command locale} or @{command class} may involve a + @{keyword "begin"} that needs to be matched by @{command (local) + "end"}, according to the usual rules for nested blocks. + + \end{description} +*} + + +section {* Local theory targets \label{sec:target} *} + +text {* + \begin{matharray}{rcll} + @{command_def "context"} & : & @{text "theory \ local_theory"} \\ + @{command_def (local) "end"} & : & @{text "local_theory \ theory"} \\ + \end{matharray} + + A local theory target is a context managed separately within the + enclosing theory. Contexts may introduce parameters (fixed + variables) and assumptions (hypotheses). Definitions and theorems + depending on the context may be added incrementally later on. + + \emph{Named contexts} refer to locales (cf.\ \secref{sec:locale}) or + type classes (cf.\ \secref{sec:class}); the name ``@{text "-"}'' + signifies the global theory context. + + \emph{Unnamed contexts} may introduce additional parameters and + assumptions, and results produced in the context are generalized + accordingly. Such auxiliary contexts may be nested within other + targets, like @{command "locale"}, @{command "class"}, @{command + "instantiation"}, @{command "overloading"}. + + @{rail " + @@{command context} @{syntax nameref} @'begin' + ; + @@{command context} @{syntax_ref \"includes\"}? (@{syntax context_elem} * ) @'begin' + ; + @{syntax_def target}: '(' @'in' @{syntax nameref} ')' + "} + + \begin{description} + + \item @{command "context"}~@{text "c \"} opens a named + context, by recommencing an existing locale or class @{text c}. + Note that locale and class definitions allow to include the + @{keyword "begin"} keyword as well, in order to continue the local + theory immediately after the initial specification. + + \item @{command "context"}~@{text "bundles elements \"} opens + an unnamed context, by extending the enclosing global or local + theory target by the given declaration bundles (\secref{sec:bundle}) + and context elements (@{text "\"}, @{text "\"} + etc.). This means any results stemming from definitions and proofs + in the extended context will be exported into the enclosing target + by lifting over extra parameters and premises. + + \item @{command (local) "end"} concludes the current local theory, + according to the nesting of contexts. Note that a global @{command + (global) "end"} has a different meaning: it concludes the theory + itself (\secref{sec:begin-thy}). + + \item @{text "("}@{keyword_def "in"}~@{text "c)"} given after any + local theory command specifies an immediate target, e.g.\ + ``@{command "definition"}~@{text "(\ c) \"}'' or ``@{command + "theorem"}~@{text "(\ c) \"}''. This works both in a local or + global theory context; the current target context will be suspended + for this command only. Note that ``@{text "(\ -)"}'' will + always produce a global result independently of the current target + context. + + \end{description} + + The exact meaning of results produced within a local theory context + depends on the underlying target infrastructure (locale, type class + etc.). The general idea is as follows, considering a context named + @{text c} with parameter @{text x} and assumption @{text "A[x]"}. + + Definitions are exported by introducing a global version with + additional arguments; a syntactic abbreviation links the long form + with the abstract version of the target context. For example, + @{text "a \ t[x]"} becomes @{text "c.a ?x \ t[?x]"} at the theory + level (for arbitrary @{text "?x"}), together with a local + abbreviation @{text "c \ c.a x"} in the target context (for the + fixed parameter @{text x}). + + Theorems are exported by discharging the assumptions and + generalizing the parameters of the context. For example, @{text "a: + B[x]"} becomes @{text "c.a: A[?x] \ B[?x]"}, again for arbitrary + @{text "?x"}. + + \medskip The Isabelle/HOL library contains numerous applications of + locales and classes, e.g.\ see @{file "~~/src/HOL/Algebra"}. An + example for an unnamed auxiliary contexts is given in @{file + "~~/src/HOL/Isar_Examples/Group_Context.thy"}. *} + + +section {* Bundled declarations \label{sec:bundle} *} + +text {* + \begin{matharray}{rcl} + @{command_def "bundle"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "print_bundles"}@{text "\<^sup>*"} & : & @{text "context \ "} \\ + @{command_def "include"} & : & @{text "proof(state) \ proof(state)"} \\ + @{command_def "including"} & : & @{text "proof(prove) \ proof(prove)"} \\ + @{keyword_def "includes"} & : & syntax \\ + \end{matharray} + + The outer syntax of fact expressions (\secref{sec:syn-att}) involves + theorems and attributes, which are evaluated in the context and + applied to it. Attributes may declare theorems to the context, as + in @{text "this_rule [intro] that_rule [elim]"} for example. + Configuration options (\secref{sec:config}) are special declaration + attributes that operate on the context without a theorem, as in + @{text "[[show_types = false]]"} for example. + + Expressions of this form may be defined as \emph{bundled + declarations} in the context, and included in other situations later + on. Including declaration bundles augments a local context casually + without logical dependencies, which is in contrast to locales and + locale interpretation (\secref{sec:locale}). + + @{rail " + @@{command bundle} @{syntax target}? \\ + @{syntax name} '=' @{syntax thmrefs} (@'for' (@{syntax vars} + @'and'))? + ; + (@@{command include} | @@{command including}) (@{syntax nameref}+) + ; + @{syntax_def \"includes\"}: @'includes' (@{syntax nameref}+) + "} + + \begin{description} + + \item @{command bundle}~@{text "b = decls"} defines a bundle of + declarations in the current context. The RHS is similar to the one + of the @{command declare} command. Bundles defined in local theory + targets are subject to transformations via morphisms, when moved + into different application contexts; this works analogously to any + other local theory specification. + + \item @{command print_bundles} prints the named bundles that are + available in the current context. + + \item @{command include}~@{text "b\<^sub>1 \ b\<^sub>n"} includes the declarations + from the given bundles into the current proof body context. This is + analogous to @{command "note"} (\secref{sec:proof-facts}) with the + expanded bundles. + + \item @{command including} is similar to @{command include}, but + works in proof refinement (backward mode). This is analogous to + @{command "using"} (\secref{sec:proof-facts}) with the expanded + bundles. + + \item @{keyword "includes"}~@{text "b\<^sub>1 \ b\<^sub>n"} is similar to + @{command include}, but works in situations where a specification + context is constructed, notably for @{command context} and long + statements of @{command theorem} etc. + + \end{description} + + Here is an artificial example of bundling various configuration + options: *} + +bundle trace = [[simp_trace, blast_trace, linarith_trace, metis_trace, smt_trace]] + +lemma "x = x" + including trace by metis + + +section {* Basic specification elements *} + +text {* + \begin{matharray}{rcll} + @{command_def "axiomatization"} & : & @{text "theory \ theory"} & (axiomatic!) \\ + @{command_def "definition"} & : & @{text "local_theory \ local_theory"} \\ + @{attribute_def "defn"} & : & @{text attribute} \\ + @{command_def "abbreviation"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "print_abbrevs"}@{text "\<^sup>*"} & : & @{text "context \ "} \\ + \end{matharray} + + These specification mechanisms provide a slightly more abstract view + than the underlying primitives of @{command "consts"}, @{command + "defs"} (see \secref{sec:consts}), and @{command "axioms"} (see + \secref{sec:axms-thms}). In particular, type-inference is commonly + available, and result names need not be given. + + @{rail " + @@{command axiomatization} @{syntax \"fixes\"}? (@'where' specs)? + ; + @@{command definition} @{syntax target}? \\ + (decl @'where')? @{syntax thmdecl}? @{syntax prop} + ; + @@{command abbreviation} @{syntax target}? @{syntax mode}? \\ + (decl @'where')? @{syntax prop} + ; + + @{syntax_def \"fixes\"}: ((@{syntax name} ('::' @{syntax type})? + @{syntax mixfix}? | @{syntax vars}) + @'and') + ; + specs: (@{syntax thmdecl}? @{syntax props} + @'and') + ; + decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? + "} + + \begin{description} + + \item @{command "axiomatization"}~@{text "c\<^sub>1 \ c\<^sub>m \ \\<^sub>1 \ \\<^sub>n"} + introduces several constants simultaneously and states axiomatic + properties for these. The constants are marked as being specified + once and for all, which prevents additional specifications being + issued later on. + + Note that axiomatic specifications are only appropriate when + declaring a new logical system; axiomatic specifications are + restricted to global theory contexts. Normal applications should + only use definitional mechanisms! + + \item @{command "definition"}~@{text "c \ eq"} produces an + internal definition @{text "c \ t"} according to the specification + given as @{text eq}, which is then turned into a proven fact. The + given proposition may deviate from internal meta-level equality + according to the rewrite rules declared as @{attribute defn} by the + object-logic. This usually covers object-level equality @{text "x = + y"} and equivalence @{text "A \ B"}. End-users normally need not + change the @{attribute defn} setup. + + Definitions may be presented with explicit arguments on the LHS, as + well as additional conditions, e.g.\ @{text "f x y = t"} instead of + @{text "f \ \x y. t"} and @{text "y \ 0 \ g x y = u"} instead of an + unrestricted @{text "g \ \x y. u"}. + + \item @{command "abbreviation"}~@{text "c \ eq"} introduces a + syntactic constant which is associated with a certain term according + to the meta-level equality @{text eq}. + + Abbreviations participate in the usual type-inference process, but + are expanded before the logic ever sees them. Pretty printing of + terms involves higher-order rewriting with rules stemming from + reverted abbreviations. This needs some care to avoid overlapping + or looping syntactic replacements! + + The optional @{text mode} specification restricts output to a + particular print mode; using ``@{text input}'' here achieves the + effect of one-way abbreviations. The mode may also include an + ``@{keyword "output"}'' qualifier that affects the concrete syntax + declared for abbreviations, cf.\ @{command "syntax"} in + \secref{sec:syn-trans}. + + \item @{command "print_abbrevs"} prints all constant abbreviations + of the current context. + + \end{description} +*} + + +section {* Generic declarations *} + +text {* + \begin{matharray}{rcl} + @{command_def "declaration"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "syntax_declaration"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "declare"} & : & @{text "local_theory \ local_theory"} \\ + \end{matharray} + + Arbitrary operations on the background context may be wrapped-up as + generic declaration elements. Since the underlying concept of local + theories may be subject to later re-interpretation, there is an + additional dependency on a morphism that tells the difference of the + original declaration context wrt.\ the application context + encountered later on. A fact declaration is an important special + case: it consists of a theorem which is applied to the context by + means of an attribute. + + @{rail " + (@@{command declaration} | @@{command syntax_declaration}) + ('(' @'pervasive' ')')? \\ @{syntax target}? @{syntax text} + ; + @@{command declare} @{syntax target}? (@{syntax thmrefs} + @'and') + "} + + \begin{description} + + \item @{command "declaration"}~@{text d} adds the declaration + function @{text d} of ML type @{ML_type declaration}, to the current + local theory under construction. In later application contexts, the + function is transformed according to the morphisms being involved in + the interpretation hierarchy. + + If the @{text "(pervasive)"} option is given, the corresponding + declaration is applied to all possible contexts involved, including + the global background theory. + + \item @{command "syntax_declaration"} is similar to @{command + "declaration"}, but is meant to affect only ``syntactic'' tools by + convention (such as notation and type-checking information). + + \item @{command "declare"}~@{text thms} declares theorems to the + current local theory context. No theorem binding is involved here, + unlike @{command "theorems"} or @{command "lemmas"} (cf.\ + \secref{sec:axms-thms}), so @{command "declare"} only has the effect + of applying attributes as included in the theorem specification. + + \end{description} +*} + + +section {* Locales \label{sec:locale} *} + +text {* + Locales are parametric named local contexts, consisting of a list of + declaration elements that are modeled after the Isar proof context + commands (cf.\ \secref{sec:proof-context}). +*} + + +subsection {* Locale expressions \label{sec:locale-expr} *} + +text {* + A \emph{locale expression} denotes a structured context composed of + instances of existing locales. The context consists of a list of + instances of declaration elements from the locales. Two locale + instances are equal if they are of the same locale and the + parameters are instantiated with equivalent terms. Declaration + elements from equal instances are never repeated, thus avoiding + duplicate declarations. More precisely, declarations from instances + that are subsumed by earlier instances are omitted. + + @{rail " + @{syntax_def locale_expr}: (instance + '+') (@'for' (@{syntax \"fixes\"} + @'and'))? + ; + instance: (qualifier ':')? @{syntax nameref} (pos_insts | named_insts) + ; + qualifier: @{syntax name} ('?' | '!')? + ; + pos_insts: ('_' | @{syntax term})* + ; + named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and') + "} + + A locale instance consists of a reference to a locale and either + positional or named parameter instantiations. Identical + instantiations (that is, those that instante a parameter by itself) + may be omitted. The notation `@{text "_"}' enables to omit the + instantiation for a parameter inside a positional instantiation. + + Terms in instantiations are from the context the locale expressions + is declared in. Local names may be added to this context with the + optional @{keyword "for"} clause. This is useful for shadowing names + bound in outer contexts, and for declaring syntax. In addition, + syntax declarations from one instance are effective when parsing + subsequent instances of the same expression. + + Instances have an optional qualifier which applies to names in + declarations. Names include local definitions and theorem names. + If present, the qualifier itself is either optional + (``\texttt{?}''), which means that it may be omitted on input of the + qualified name, or mandatory (``\texttt{!}''). If neither + ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default + is used. For @{command "interpretation"} and @{command "interpret"} + the default is ``mandatory'', for @{command "locale"} and @{command + "sublocale"} the default is ``optional''. +*} + + +subsection {* Locale declarations *} + +text {* + \begin{matharray}{rcl} + @{command_def "locale"} & : & @{text "theory \ local_theory"} \\ + @{command_def "print_locale"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_locales"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{method_def intro_locales} & : & @{text method} \\ + @{method_def unfold_locales} & : & @{text method} \\ + \end{matharray} + + \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes} + \indexisarelem{defines}\indexisarelem{notes} + @{rail " + @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'? + ; + @@{command print_locale} '!'? @{syntax nameref} + ; + @{syntax_def locale}: @{syntax context_elem}+ | + @{syntax locale_expr} ('+' (@{syntax context_elem}+))? + ; + @{syntax_def context_elem}: + @'fixes' (@{syntax \"fixes\"} + @'and') | + @'constrains' (@{syntax name} '::' @{syntax type} + @'and') | + @'assumes' (@{syntax props} + @'and') | + @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') | + @'notes' (@{syntax thmdef}? @{syntax thmrefs} + @'and') + "} + + \begin{description} + + \item @{command "locale"}~@{text "loc = import + body"} defines a + new locale @{text loc} as a context consisting of a certain view of + existing locales (@{text import}) plus some additional elements + (@{text body}). Both @{text import} and @{text body} are optional; + the degenerate form @{command "locale"}~@{text loc} defines an empty + locale, which may still be useful to collect declarations of facts + later on. Type-inference on locale expressions automatically takes + care of the most general typing that the combined context elements + may acquire. + + The @{text import} consists of a structured locale expression; see + \secref{sec:proof-context} above. Its for clause defines the local + parameters of the @{text import}. In addition, locale parameters + whose instantance is omitted automatically extend the (possibly + empty) for clause: they are inserted at its beginning. This means + that these parameters may be referred to from within the expression + and also in the subsequent context elements and provides a + notational convenience for the inheritance of parameters in locale + declarations. + + The @{text body} consists of context elements. + + \begin{description} + + \item @{element "fixes"}~@{text "x :: \ (mx)"} declares a local + parameter of type @{text \} and mixfix annotation @{text mx} (both + are optional). The special syntax declaration ``@{text + "(\)"}'' means that @{text x} may be referenced + implicitly in this context. + + \item @{element "constrains"}~@{text "x :: \"} introduces a type + constraint @{text \} on the local parameter @{text x}. This + element is deprecated. The type constraint should be introduced in + the for clause or the relevant @{element "fixes"} element. + + \item @{element "assumes"}~@{text "a: \\<^sub>1 \ \\<^sub>n"} + introduces local premises, similar to @{command "assume"} within a + proof (cf.\ \secref{sec:proof-context}). + + \item @{element "defines"}~@{text "a: x \ t"} defines a previously + declared parameter. This is similar to @{command "def"} within a + proof (cf.\ \secref{sec:proof-context}), but @{element "defines"} + takes an equational proposition instead of variable-term pair. The + left-hand side of the equation may have additional arguments, e.g.\ + ``@{element "defines"}~@{text "f x\<^sub>1 \ x\<^sub>n \ t"}''. + + \item @{element "notes"}~@{text "a = b\<^sub>1 \ b\<^sub>n"} + reconsiders facts within a local context. Most notably, this may + include arbitrary declarations in any attribute specifications + included here, e.g.\ a local @{attribute simp} rule. + + The initial @{text import} specification of a locale expression + maintains a dynamic relation to the locales being referenced + (benefiting from any later fact declarations in the obvious manner). + + \end{description} + + Note that ``@{text "(\ p\<^sub>1 \ p\<^sub>n)"}'' patterns given + in the syntax of @{element "assumes"} and @{element "defines"} above + are illegal in locale definitions. In the long goal format of + \secref{sec:goals}, term bindings may be included as expected, + though. + + \medskip Locale specifications are ``closed up'' by + turning the given text into a predicate definition @{text + loc_axioms} and deriving the original assumptions as local lemmas + (modulo local definitions). The predicate statement covers only the + newly specified assumptions, omitting the content of included locale + expressions. The full cumulative view is only provided on export, + involving another predicate @{text loc} that refers to the complete + specification text. + + In any case, the predicate arguments are those locale parameters + that actually occur in the respective piece of text. Also note that + these predicates operate at the meta-level in theory, but the locale + packages attempts to internalize statements according to the + object-logic setup (e.g.\ replacing @{text \} by @{text \}, and + @{text "\"} by @{text "\"} in HOL; see also + \secref{sec:object-logic}). Separate introduction rules @{text + loc_axioms.intro} and @{text loc.intro} are provided as well. + + \item @{command "print_locale"}~@{text "locale"} prints the + contents of the named locale. The command omits @{element "notes"} + elements by default. Use @{command "print_locale"}@{text "!"} to + have them included. + + \item @{command "print_locales"} prints the names of all locales + of the current theory. + + \item @{method intro_locales} and @{method unfold_locales} + repeatedly expand all introduction rules of locale predicates of the + theory. While @{method intro_locales} only applies the @{text + loc.intro} introduction rules and therefore does not decend to + assumptions, @{method unfold_locales} is more aggressive and applies + @{text loc_axioms.intro} as well. Both methods are aware of locale + specifications entailed by the context, both from target statements, + and from interpretations (see below). New goals that are entailed + by the current context are discharged automatically. + + \end{description} +*} + + +subsection {* Locale interpretation *} + +text {* + \begin{matharray}{rcl} + @{command_def "interpretation"} & : & @{text "theory \ proof(prove)"} \\ + @{command_def "interpret"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ + @{command_def "sublocale"} & : & @{text "theory \ proof(prove)"} \\ + @{command_def "print_dependencies"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "print_interps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + \end{matharray} + + Locale expressions may be instantiated, and the instantiated facts + added to the current context. This requires a proof of the + instantiated specification and is called \emph{locale + interpretation}. Interpretation is possible in locales (command + @{command "sublocale"}), theories (command @{command + "interpretation"}) and also within proof bodies (command @{command + "interpret"}). + + @{rail " + @@{command interpretation} @{syntax locale_expr} equations? + ; + @@{command interpret} @{syntax locale_expr} equations? + ; + @@{command sublocale} @{syntax nameref} ('<' | '\') @{syntax locale_expr} \\ + equations? + ; + @@{command print_dependencies} '!'? @{syntax locale_expr} + ; + @@{command print_interps} @{syntax nameref} + ; + + equations: @'where' (@{syntax thmdecl}? @{syntax prop} + @'and') + "} + + \begin{description} + + \item @{command "interpretation"}~@{text "expr \ eqns"} + interprets @{text expr} in the theory. The command generates proof + obligations for the instantiated specifications (assumes and defines + elements). Once these are discharged by the user, instantiated + facts are added to the theory in a post-processing phase. + + Free variables in the interpreted expression are allowed. They are + turned into schematic variables in the generated declarations. In + order to use a free variable whose name is already bound in the + context --- for example, because a constant of that name exists --- + it may be added to the @{keyword "for"} clause. + + Additional equations, which are unfolded during + post-processing, may be given after the keyword @{keyword "where"}. + This is useful for interpreting concepts introduced through + definitions. The equations must be proved. + + The command is aware of interpretations already active in the + theory, but does not simplify the goal automatically. In order to + simplify the proof obligations use methods @{method intro_locales} + or @{method unfold_locales}. Post-processing is not applied to + facts of interpretations that are already active. This avoids + duplication of interpreted facts, in particular. Note that, in the + case of a locale with import, parts of the interpretation may + already be active. The command will only process facts for new + parts. + + Adding facts to locales has the effect of adding interpreted facts + to the theory for all interpretations as well. That is, + interpretations dynamically participate in any facts added to + locales. Note that if a theory inherits additional facts for a + locale through one parent and an interpretation of that locale + through another parent, the additional facts will not be + interpreted. + + \item @{command "interpret"}~@{text "expr \ eqns"} interprets + @{text expr} in the proof context and is otherwise similar to + interpretation in theories. Note that rewrite rules given to + @{command "interpret"} after the @{keyword "where"} keyword should be + explicitly universally quantified. + + \item @{command "sublocale"}~@{text "name \ expr \ + eqns"} + interprets @{text expr} in the locale @{text name}. A proof that + the specification of @{text name} implies the specification of + @{text expr} is required. As in the localized version of the + theorem command, the proof is in the context of @{text name}. After + the proof obligation has been discharged, the facts of @{text expr} + become part of locale @{text name} as \emph{derived} context + elements and are available when the context @{text name} is + subsequently entered. Note that, like import, this is dynamic: + facts added to a locale part of @{text expr} after interpretation + become also available in @{text name}. + + Only specification fragments of @{text expr} that are not already + part of @{text name} (be it imported, derived or a derived fragment + of the import) are considered in this process. This enables + circular interpretations provided that no infinite chains are + generated in the locale hierarchy. + + If interpretations of @{text name} exist in the current theory, the + command adds interpretations for @{text expr} as well, with the same + qualifier, although only for fragments of @{text expr} that are not + interpreted in the theory already. + + Equations given after @{keyword "where"} amend the morphism through + which @{text expr} is interpreted. This enables to map definitions + from the interpreted locales to entities of @{text name}. This + feature is experimental. + + \item @{command "print_dependencies"}~@{text "expr"} is useful for + understanding the effect of an interpretation of @{text "expr"}. It + lists all locale instances for which interpretations would be added + to the current context. Variant @{command + "print_dependencies"}@{text "!"} prints all locale instances that + would be considered for interpretation, and would be interpreted in + an empty context (that is, without interpretations). + + \item @{command "print_interps"}~@{text "locale"} lists all + interpretations of @{text "locale"} in the current theory or proof + context, including those due to a combination of a @{command + "interpretation"} or @{command "interpret"} and one or several + @{command "sublocale"} declarations. + + \end{description} + + \begin{warn} + Since attributes are applied to interpreted theorems, + interpretation may modify the context of common proof tools, e.g.\ + the Simplifier or Classical Reasoner. As the behavior of such + tools is \emph{not} stable under interpretation morphisms, manual + declarations might have to be added to the target context of the + interpretation to revert such declarations. + \end{warn} + + \begin{warn} + An interpretation in a theory or proof context may subsume previous + interpretations. This happens if the same specification fragment + is interpreted twice and the instantiation of the second + interpretation is more general than the interpretation of the + first. The locale package does not attempt to remove subsumed + interpretations. + \end{warn} +*} + + +section {* Classes \label{sec:class} *} + +text {* + \begin{matharray}{rcl} + @{command_def "class"} & : & @{text "theory \ local_theory"} \\ + @{command_def "instantiation"} & : & @{text "theory \ local_theory"} \\ + @{command_def "instance"} & : & @{text "local_theory \ local_theory"} \\ + @{command "instance"} & : & @{text "theory \ proof(prove)"} \\ + @{command_def "subclass"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "print_classes"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ + @{method_def intro_classes} & : & @{text method} \\ + \end{matharray} + + A class is a particular locale with \emph{exactly one} type variable + @{text \}. Beyond the underlying locale, a corresponding type class + is established which is interpreted logically as axiomatic type + class \cite{Wenzel:1997:TPHOL} whose logical content are the + assumptions of the locale. Thus, classes provide the full + generality of locales combined with the commodity of type classes + (notably type-inference). See \cite{isabelle-classes} for a short + tutorial. + + @{rail " + @@{command class} class_spec @'begin'? + ; + class_spec: @{syntax name} '=' + ((@{syntax nameref} '+' (@{syntax context_elem}+)) | + @{syntax nameref} | (@{syntax context_elem}+)) + ; + @@{command instantiation} (@{syntax nameref} + @'and') '::' @{syntax arity} @'begin' + ; + @@{command instance} (() | (@{syntax nameref} + @'and') '::' @{syntax arity} | + @{syntax nameref} ('<' | '\') @{syntax nameref} ) + ; + @@{command subclass} @{syntax target}? @{syntax nameref} + "} + + \begin{description} + + \item @{command "class"}~@{text "c = superclasses + body"} defines + a new class @{text c}, inheriting from @{text superclasses}. This + introduces a locale @{text c} with import of all locales @{text + superclasses}. + + Any @{element "fixes"} in @{text body} are lifted to the global + theory level (\emph{class operations} @{text "f\<^sub>1, \, + f\<^sub>n"} of class @{text c}), mapping the local type parameter + @{text \} to a schematic type variable @{text "?\ :: c"}. + + Likewise, @{element "assumes"} in @{text body} are also lifted, + mapping each local parameter @{text "f :: \[\]"} to its + corresponding global constant @{text "f :: \[?\ :: c]"}. The + corresponding introduction rule is provided as @{text + c_class_axioms.intro}. This rule should be rarely needed directly + --- the @{method intro_classes} method takes care of the details of + class membership proofs. + + \item @{command "instantiation"}~@{text "t :: (s\<^sub>1, \, s\<^sub>n)s + \"} opens a theory target (cf.\ \secref{sec:target}) which + allows to specify class operations @{text "f\<^sub>1, \, f\<^sub>n"} corresponding + to sort @{text s} at the particular type instance @{text "(\\<^sub>1 :: s\<^sub>1, + \, \\<^sub>n :: s\<^sub>n) t"}. A plain @{command "instance"} command in the + target body poses a goal stating these type arities. The target is + concluded by an @{command_ref (local) "end"} command. + + Note that a list of simultaneous type constructors may be given; + this corresponds nicely to mutually recursive type definitions, e.g.\ + in Isabelle/HOL. + + \item @{command "instance"} in an instantiation target body sets + up a goal stating the type arities claimed at the opening @{command + "instantiation"}. The proof would usually proceed by @{method + intro_classes}, and then establish the characteristic theorems of + the type classes involved. After finishing the proof, the + background theory will be augmented by the proven type arities. + + On the theory level, @{command "instance"}~@{text "t :: (s\<^sub>1, \, + s\<^sub>n)s"} provides a convenient way to instantiate a type class with no + need to specify operations: one can continue with the + instantiation proof immediately. + + \item @{command "subclass"}~@{text c} in a class context for class + @{text d} sets up a goal stating that class @{text c} is logically + contained in class @{text d}. After finishing the proof, class + @{text d} is proven to be subclass @{text c} and the locale @{text + c} is interpreted into @{text d} simultaneously. + + A weakend form of this is available through a further variant of + @{command instance}: @{command instance}~@{text "c\<^sub>1 \ c\<^sub>2"} opens + a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference + to the underlying locales; this is useful if the properties to prove + the logical connection are not sufficent on the locale level but on + the theory level. + + \item @{command "print_classes"} prints all classes in the current + theory. + + \item @{command "class_deps"} visualizes all classes and their + subclass relations as a Hasse diagram. + + \item @{method intro_classes} repeatedly expands all class + introduction rules of this theory. Note that this method usually + needs not be named explicitly, as it is already included in the + default proof step (e.g.\ of @{command "proof"}). In particular, + instantiation of trivial (syntactic) classes may be performed by a + single ``@{command ".."}'' proof step. + + \end{description} +*} + + +subsection {* The class target *} + +text {* + %FIXME check + + A named context may refer to a locale (cf.\ \secref{sec:target}). + If this locale is also a class @{text c}, apart from the common + locale target behaviour the following happens. + + \begin{itemize} + + \item Local constant declarations @{text "g[\]"} referring to the + local type parameter @{text \} and local parameters @{text "f[\]"} + are accompanied by theory-level constants @{text "g[?\ :: c]"} + referring to theory-level class operations @{text "f[?\ :: c]"}. + + \item Local theorem bindings are lifted as are assumptions. + + \item Local syntax refers to local operations @{text "g[\]"} and + global operations @{text "g[?\ :: c]"} uniformly. Type inference + resolves ambiguities. In rare cases, manual type annotations are + needed. + + \end{itemize} +*} + + +subsection {* Co-regularity of type classes and arities *} + +text {* The class relation together with the collection of + type-constructor arities must obey the principle of + \emph{co-regularity} as defined below. + + \medskip For the subsequent formulation of co-regularity we assume + that the class relation is closed by transitivity and reflexivity. + Moreover the collection of arities @{text "t :: (\<^vec>s)c"} is + completed such that @{text "t :: (\<^vec>s)c"} and @{text "c \ c'"} + implies @{text "t :: (\<^vec>s)c'"} for all such declarations. + + Treating sorts as finite sets of classes (meaning the intersection), + the class relation @{text "c\<^sub>1 \ c\<^sub>2"} is extended to sorts as + follows: + \[ + @{text "s\<^sub>1 \ s\<^sub>2 \ \c\<^sub>2 \ s\<^sub>2. \c\<^sub>1 \ s\<^sub>1. c\<^sub>1 \ c\<^sub>2"} + \] + + This relation on sorts is further extended to tuples of sorts (of + the same length) in the component-wise way. + + \smallskip Co-regularity of the class relation together with the + arities relation means: + \[ + @{text "t :: (\<^vec>s\<^sub>1)c\<^sub>1 \ t :: (\<^vec>s\<^sub>2)c\<^sub>2 \ c\<^sub>1 \ c\<^sub>2 \ \<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2"} + \] + \noindent for all such arities. In other words, whenever the result + classes of some type-constructor arities are related, then the + argument sorts need to be related in the same way. + + \medskip Co-regularity is a very fundamental property of the + order-sorted algebra of types. For example, it entails principle + types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}. +*} + + +section {* Unrestricted overloading *} + +text {* + \begin{matharray}{rcl} + @{command_def "overloading"} & : & @{text "theory \ local_theory"} \\ + \end{matharray} + + Isabelle/Pure's definitional schemes support certain forms of + overloading (see \secref{sec:consts}). Overloading means that a + constant being declared as @{text "c :: \ decl"} may be + defined separately on type instances + @{text "c :: (\\<^sub>1, \, \\<^sub>n) t decl"} + for each type constructor @{text t}. At most occassions + overloading will be used in a Haskell-like fashion together with + type classes by means of @{command "instantiation"} (see + \secref{sec:class}). Sometimes low-level overloading is desirable. + The @{command "overloading"} target provides a convenient view for + end-users. + + @{rail " + @@{command overloading} ( spec + ) @'begin' + ; + spec: @{syntax name} ( '==' | '\' ) @{syntax term} ( '(' @'unchecked' ')' )? + "} + + \begin{description} + + \item @{command "overloading"}~@{text "x\<^sub>1 \ c\<^sub>1 :: \\<^sub>1 \ \ x\<^sub>n \ c\<^sub>n :: \\<^sub>n \"} + opens a theory target (cf.\ \secref{sec:target}) which allows to + specify constants with overloaded definitions. These are identified + by an explicitly given mapping from variable names @{text "x\<^sub>i"} to + constants @{text "c\<^sub>i"} at particular type instances. The + definitions themselves are established using common specification + tools, using the names @{text "x\<^sub>i"} as reference to the + corresponding constants. The target is concluded by @{command + (local) "end"}. + + A @{text "(unchecked)"} option disables global dependency checks for + the corresponding definition, which is occasionally useful for + exotic overloading (see \secref{sec:consts} for a precise description). + It is at the discretion of the user to avoid + malformed theory specifications! + + \end{description} +*} + + +section {* Incorporating ML code \label{sec:ML} *} + +text {* + \begin{matharray}{rcl} + @{command_def "ML_file"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "ML"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "ML_prf"} & : & @{text "proof \ proof"} \\ + @{command_def "ML_val"} & : & @{text "any \"} \\ + @{command_def "ML_command"} & : & @{text "any \"} \\ + @{command_def "setup"} & : & @{text "theory \ theory"} \\ + @{command_def "local_setup"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "attribute_setup"} & : & @{text "theory \ theory"} \\ + \end{matharray} + + @{rail " + @@{command ML_file} @{syntax name} + ; + (@@{command ML} | @@{command ML_prf} | @@{command ML_val} | + @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text} + ; + @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}? + "} + + \begin{description} + + \item @{command "ML_file"}~@{text "name"} reads and evaluates the + given ML file. The current theory context is passed down to the ML + toplevel and may be modified, using @{ML "Context.>>"} or derived ML + commands. Top-level ML bindings are stored within the (global or + local) theory context. + + \item @{command "ML"}~@{text "text"} is similar to @{command + "ML_file"}, but evaluates directly the given @{text "text"}. + Top-level ML bindings are stored within the (global or local) theory + context. + + \item @{command "ML_prf"} is analogous to @{command "ML"} but works + within a proof context. Top-level ML bindings are stored within the + proof context in a purely sequential fashion, disregarding the + nested proof structure. ML bindings introduced by @{command + "ML_prf"} are discarded at the end of the proof. + + \item @{command "ML_val"} and @{command "ML_command"} are diagnostic + versions of @{command "ML"}, which means that the context may not be + updated. @{command "ML_val"} echos the bindings produced at the ML + toplevel, but @{command "ML_command"} is silent. + + \item @{command "setup"}~@{text "text"} changes the current theory + context by applying @{text "text"}, which refers to an ML expression + of type @{ML_type "theory -> theory"}. This enables to initialize + any object-logic specific tools and packages written in ML, for + example. + + \item @{command "local_setup"} is similar to @{command "setup"} for + a local theory context, and an ML expression of type @{ML_type + "local_theory -> local_theory"}. This allows to + invoke local theory specification packages without going through + concrete outer syntax, for example. + + \item @{command "attribute_setup"}~@{text "name = text description"} + defines an attribute in the current theory. The given @{text + "text"} has to be an ML expression of type + @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in + structure @{ML_struct Args} and @{ML_struct Attrib}. + + In principle, attributes can operate both on a given theorem and the + implicit context, although in practice only one is modified and the + other serves as parameter. Here are examples for these two cases: + + \end{description} +*} + + attribute_setup my_rule = {* + Attrib.thms >> (fn ths => + Thm.rule_attribute + (fn context: Context.generic => fn th: thm => + let val th' = th OF ths + in th' end)) *} + + attribute_setup my_declaration = {* + Attrib.thms >> (fn ths => + Thm.declaration_attribute + (fn th: thm => fn context: Context.generic => + let val context' = context + in context' end)) *} + + +section {* Primitive specification elements *} + +subsection {* Type classes and sorts \label{sec:classes} *} + +text {* + \begin{matharray}{rcll} + @{command_def "classes"} & : & @{text "theory \ theory"} \\ + @{command_def "classrel"} & : & @{text "theory \ theory"} & (axiomatic!) \\ + @{command_def "default_sort"} & : & @{text "local_theory \ local_theory"} + \end{matharray} + + @{rail " + @@{command classes} (@{syntax classdecl} +) + ; + @@{command classrel} (@{syntax nameref} ('<' | '\') @{syntax nameref} + @'and') + ; + @@{command default_sort} @{syntax sort} + "} + + \begin{description} + + \item @{command "classes"}~@{text "c \ c\<^sub>1, \, c\<^sub>n"} declares class + @{text c} to be a subclass of existing classes @{text "c\<^sub>1, \, c\<^sub>n"}. + Isabelle implicitly maintains the transitive closure of the class + hierarchy. Cyclic class structures are not permitted. + + \item @{command "classrel"}~@{text "c\<^sub>1 \ c\<^sub>2"} states subclass + relations between existing classes @{text "c\<^sub>1"} and @{text "c\<^sub>2"}. + This is done axiomatically! The @{command_ref "subclass"} and + @{command_ref "instance"} commands (see \secref{sec:class}) provide + a way to introduce proven class relations. + + \item @{command "default_sort"}~@{text s} makes sort @{text s} the + new default sort for any type variable that is given explicitly in + the text, but lacks a sort constraint (wrt.\ the current context). + Type variables generated by type inference are not affected. + + Usually the default sort is only changed when defining a new + object-logic. For example, the default sort in Isabelle/HOL is + @{class type}, the class of all HOL types. + + When merging theories, the default sorts of the parents are + logically intersected, i.e.\ the representations as lists of classes + are joined. + + \end{description} +*} + + +subsection {* Types and type abbreviations \label{sec:types-pure} *} + +text {* + \begin{matharray}{rcll} + @{command_def "type_synonym"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "typedecl"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "arities"} & : & @{text "theory \ theory"} & (axiomatic!) \\ + \end{matharray} + + @{rail " + @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?) + ; + @@{command typedecl} @{syntax typespec} @{syntax mixfix}? + ; + @@{command arities} (@{syntax nameref} '::' @{syntax arity} +) + "} + + \begin{description} + + \item @{command "type_synonym"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = \"} + introduces a \emph{type synonym} @{text "(\\<^sub>1, \, \\<^sub>n) t"} for the + existing type @{text "\"}. Unlike actual type definitions, as are + available in Isabelle/HOL for example, type synonyms are merely + syntactic abbreviations without any logical significance. + Internally, type synonyms are fully expanded. + + \item @{command "typedecl"}~@{text "(\\<^sub>1, \, \\<^sub>n) t"} declares a new + type constructor @{text t}. If the object-logic defines a base sort + @{text s}, then the constructor is declared to operate on that, via + the axiomatic specification @{command arities}~@{text "t :: (s, \, + s)s"}. + + \item @{command "arities"}~@{text "t :: (s\<^sub>1, \, s\<^sub>n)s"} augments + Isabelle's order-sorted signature of types by new type constructor + arities. This is done axiomatically! The @{command_ref "instantiation"} + target (see \secref{sec:class}) provides a way to introduce + proven type arities. + + \end{description} +*} + + +subsection {* Constants and definitions \label{sec:consts} *} + +text {* + \begin{matharray}{rcl} + @{command_def "consts"} & : & @{text "theory \ theory"} \\ + @{command_def "defs"} & : & @{text "theory \ theory"} \\ + \end{matharray} + + Definitions essentially express abbreviations within the logic. The + simplest form of a definition is @{text "c :: \ \ t"}, where @{text + c} is a newly declared constant. Isabelle also allows derived forms + where the arguments of @{text c} appear on the left, abbreviating a + prefix of @{text \}-abstractions, e.g.\ @{text "c \ \x y. t"} may be + written more conveniently as @{text "c x y \ t"}. Moreover, + definitions may be weakened by adding arbitrary pre-conditions: + @{text "A \ c x y \ t"}. + + \medskip The built-in well-formedness conditions for definitional + specifications are: + + \begin{itemize} + + \item Arguments (on the left-hand side) must be distinct variables. + + \item All variables on the right-hand side must also appear on the + left-hand side. + + \item All type variables on the right-hand side must also appear on + the left-hand side; this prohibits @{text "0 :: nat \ length ([] :: + \ list)"} for example. + + \item The definition must not be recursive. Most object-logics + provide definitional principles that can be used to express + recursion safely. + + \end{itemize} + + The right-hand side of overloaded definitions may mention overloaded constants + recursively at type instances corresponding to the immediate + argument types @{text "\\<^sub>1, \, \\<^sub>n"}. Incomplete + specification patterns impose global constraints on all occurrences, + e.g.\ @{text "d :: \ \ \"} on the left-hand side means that all + corresponding occurrences on some right-hand side need to be an + instance of this, general @{text "d :: \ \ \"} will be disallowed. + + @{rail " + @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +) + ; + @@{command defs} opt? (@{syntax axmdecl} @{syntax prop} +) + ; + opt: '(' @'unchecked'? @'overloaded'? ')' + "} + + \begin{description} + + \item @{command "consts"}~@{text "c :: \"} declares constant @{text + c} to have any instance of type scheme @{text \}. The optional + mixfix annotations may attach concrete syntax to the constants + declared. + + \item @{command "defs"}~@{text "name: eqn"} introduces @{text eqn} + as a definitional axiom for some existing constant. + + The @{text "(unchecked)"} option disables global dependency checks + for this definition, which is occasionally useful for exotic + overloading. It is at the discretion of the user to avoid malformed + theory specifications! + + The @{text "(overloaded)"} option declares definitions to be + potentially overloaded. Unless this option is given, a warning + message would be issued for any definitional equation with a more + special type than that of the corresponding constant declaration. + + \end{description} +*} + + +section {* Axioms and theorems \label{sec:axms-thms} *} + +text {* + \begin{matharray}{rcll} + @{command_def "axioms"} & : & @{text "theory \ theory"} & (axiomatic!) \\ + @{command_def "lemmas"} & : & @{text "local_theory \ local_theory"} \\ + @{command_def "theorems"} & : & @{text "local_theory \ local_theory"} \\ + \end{matharray} + + @{rail " + @@{command axioms} (@{syntax axmdecl} @{syntax prop} +) + ; + (@@{command lemmas} | @@{command theorems}) @{syntax target}? \\ + (@{syntax thmdef}? @{syntax thmrefs} + @'and') + (@'for' (@{syntax vars} + @'and'))? + "} + + \begin{description} + + \item @{command "axioms"}~@{text "a: \"} introduces arbitrary + statements as axioms of the meta-logic. In fact, axioms are + ``axiomatic theorems'', and may be referred later just as any other + theorem. + + Axioms are usually only introduced when declaring new logical + systems. Everyday work is typically done the hard way, with proper + definitions and proven theorems. + + \item @{command "lemmas"}~@{text "a = b\<^sub>1 \ b\<^sub>n"}~@{keyword_def + "for"}~@{text "x\<^sub>1 \ x\<^sub>m"} evaluates given facts (with attributes) in + the current context, which may be augmented by local variables. + Results are standardized before being stored, i.e.\ schematic + variables are renamed to enforce index @{text "0"} uniformly. + + \item @{command "theorems"} is the same as @{command "lemmas"}, but + marks the result as a different kind of facts. + + \end{description} +*} + + +section {* Oracles *} + +text {* + \begin{matharray}{rcll} + @{command_def "oracle"} & : & @{text "theory \ theory"} & (axiomatic!) \\ + \end{matharray} + + Oracles allow Isabelle to take advantage of external reasoners such + as arithmetic decision procedures, model checkers, fast tautology + checkers or computer algebra systems. Invoked as an oracle, an + external reasoner can create arbitrary Isabelle theorems. + + It is the responsibility of the user to ensure that the external + reasoner is as trustworthy as the application requires. Another + typical source of errors is the linkup between Isabelle and the + external tool, not just its concrete implementation, but also the + required translation between two different logical environments. + + Isabelle merely guarantees well-formedness of the propositions being + asserted, and records within the internal derivation object how + presumed theorems depend on unproven suppositions. + + @{rail " + @@{command oracle} @{syntax name} '=' @{syntax text} + "} + + \begin{description} + + \item @{command "oracle"}~@{text "name = text"} turns the given ML + expression @{text "text"} of type @{ML_text "'a -> cterm"} into an + ML function of type @{ML_text "'a -> thm"}, which is bound to the + global identifier @{ML_text name}. This acts like an infinitary + specification of axioms! Invoking the oracle only works within the + scope of the resulting theory. + + \end{description} + + See @{file "~~/src/HOL/ex/Iff_Oracle.thy"} for a worked example of + defining a new primitive rule as oracle, and turning it into a proof + method. +*} + + +section {* Name spaces *} + +text {* + \begin{matharray}{rcl} + @{command_def "hide_class"} & : & @{text "theory \ theory"} \\ + @{command_def "hide_type"} & : & @{text "theory \ theory"} \\ + @{command_def "hide_const"} & : & @{text "theory \ theory"} \\ + @{command_def "hide_fact"} & : & @{text "theory \ theory"} \\ + \end{matharray} + + @{rail " + ( @{command hide_class} | @{command hide_type} | + @{command hide_const} | @{command hide_fact} ) ('(' @'open' ')')? (@{syntax nameref} + ) + "} + + Isabelle organizes any kind of name declarations (of types, + constants, theorems etc.) by separate hierarchically structured name + spaces. Normally the user does not have to control the behavior of + name spaces by hand, yet the following commands provide some way to + do so. + + \begin{description} + + \item @{command "hide_class"}~@{text names} fully removes class + declarations from a given name space; with the @{text "(open)"} + option, only the base name is hidden. + + Note that hiding name space accesses has no impact on logical + declarations --- they remain valid internally. Entities that are no + longer accessible to the user are printed with the special qualifier + ``@{text "??"}'' prefixed to the full internal name. + + \item @{command "hide_type"}, @{command "hide_const"}, and @{command + "hide_fact"} are similar to @{command "hide_class"}, but hide types, + constants, and facts, respectively. + + \end{description} +*} + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Symbols.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Symbols.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,46 @@ +theory Symbols +imports Base Main +begin + +chapter {* Predefined Isabelle symbols \label{app:symbols} *} + +text {* + Isabelle supports an infinite number of non-ASCII symbols, which are + represented in source text as @{verbatim "\\"}@{verbatim "<"}@{text + name}@{verbatim ">"} (where @{text name} may be any identifier). It + is left to front-end tools how to present these symbols to the user. + The collection of predefined standard symbols given below is + available by default for Isabelle document output, due to + appropriate definitions of @{verbatim "\\"}@{verbatim isasym}@{text + name} for each @{verbatim "\\"}@{verbatim "<"}@{text name}@{verbatim + ">"} in the @{verbatim isabellesym.sty} file. Most of these symbols + are displayed properly in Proof~General and Isabelle/jEdit. + + Moreover, any single symbol (or ASCII character) may be prefixed by + @{verbatim "\\"}@{verbatim "<^sup>"}, for superscript and @{verbatim + "\\"}@{verbatim "<^sub>"}, for subscript, such as @{verbatim + "A\\"}@{verbatim "<^sup>\"}, for @{text "A\<^sup>\"} the alternative + versions @{verbatim "\\"}@{verbatim "<^isub>"} and @{verbatim + "\\"}@{verbatim "<^isup>"} are considered as quasi letters and may + be used within identifiers. Sub- and superscripts that span a + region of text are marked up with @{verbatim "\\"}@{verbatim + "<^bsub>"}@{text "\"}@{verbatim "\\"}@{verbatim "<^esub>"}, and + @{verbatim "\\"}@{verbatim "<^bsup>"}@{text "\"}@{verbatim + "\\"}@{verbatim "<^esup>"} respectively. Furthermore, all ASCII + characters and most other symbols may be printed in bold by + prefixing @{verbatim "\\"}@{verbatim "<^bold>"} such as @{verbatim + "\\"}@{verbatim "<^bold>\\"}@{verbatim ""} for @{text + "\<^bold>\"}. Note that @{verbatim "\\"}@{verbatim "<^bold>"}, may + \emph{not} be combined with sub- or superscripts for single symbols. + + Further details of Isabelle document preparation are covered in + \chref{ch:document-prep}. + + \begin{center} + \begin{isabellebody} + \input{syms} + \end{isabellebody} + \end{center} +*} + +end \ No newline at end of file diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Synopsis.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/Synopsis.thy Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,1108 @@ +theory Synopsis +imports Base Main +begin + +chapter {* Synopsis *} + +section {* Notepad *} + +text {* + An Isar proof body serves as mathematical notepad to compose logical + content, consisting of types, terms, facts. +*} + + +subsection {* Types and terms *} + +notepad +begin + txt {* Locally fixed entities: *} + fix x -- {* local constant, without any type information yet *} + fix x :: 'a -- {* variant with explicit type-constraint for subsequent use*} + + fix a b + assume "a = b" -- {* type assignment at first occurrence in concrete term *} + + txt {* Definitions (non-polymorphic): *} + def x \ "t::'a" + + txt {* Abbreviations (polymorphic): *} + let ?f = "\x. x" + term "?f ?f" + + txt {* Notation: *} + write x ("***") +end + + +subsection {* Facts *} + +text {* + A fact is a simultaneous list of theorems. +*} + + +subsubsection {* Producing facts *} + +notepad +begin + + txt {* Via assumption (``lambda''): *} + assume a: A + + txt {* Via proof (``let''): *} + have b: B sorry + + txt {* Via abbreviation (``let''): *} + note c = a b + +end + + +subsubsection {* Referencing facts *} + +notepad +begin + txt {* Via explicit name: *} + assume a: A + note a + + txt {* Via implicit name: *} + assume A + note this + + txt {* Via literal proposition (unification with results from the proof text): *} + assume A + note `A` + + assume "\x. B x" + note `B a` + note `B b` +end + + +subsubsection {* Manipulating facts *} + +notepad +begin + txt {* Instantiation: *} + assume a: "\x. B x" + note a + note a [of b] + note a [where x = b] + + txt {* Backchaining: *} + assume 1: A + assume 2: "A \ C" + note 2 [OF 1] + note 1 [THEN 2] + + txt {* Symmetric results: *} + assume "x = y" + note this [symmetric] + + assume "x \ y" + note this [symmetric] + + txt {* Adhoc-simplification (take care!): *} + assume "P ([] @ xs)" + note this [simplified] +end + + +subsubsection {* Projections *} + +text {* + Isar facts consist of multiple theorems. There is notation to project + interval ranges. +*} + +notepad +begin + assume stuff: A B C D + note stuff(1) + note stuff(2-3) + note stuff(2-) +end + + +subsubsection {* Naming conventions *} + +text {* + \begin{itemize} + + \item Lower-case identifiers are usually preferred. + + \item Facts can be named after the main term within the proposition. + + \item Facts should \emph{not} be named after the command that + introduced them (@{command "assume"}, @{command "have"}). This is + misleading and hard to maintain. + + \item Natural numbers can be used as ``meaningless'' names (more + appropriate than @{text "a1"}, @{text "a2"} etc.) + + \item Symbolic identifiers are supported (e.g. @{text "*"}, @{text + "**"}, @{text "***"}). + + \end{itemize} +*} + + +subsection {* Block structure *} + +text {* + The formal notepad is block structured. The fact produced by the last + entry of a block is exported into the outer context. +*} + +notepad +begin + { + have a: A sorry + have b: B sorry + note a b + } + note this + note `A` + note `B` +end + +text {* Explicit blocks as well as implicit blocks of nested goal + statements (e.g.\ @{command have}) automatically introduce one extra + pair of parentheses in reserve. The @{command next} command allows + to ``jump'' between these sub-blocks. *} + +notepad +begin + + { + have a: A sorry + next + have b: B + proof - + show B sorry + next + have c: C sorry + next + have d: D sorry + qed + } + + txt {* Alternative version with explicit parentheses everywhere: *} + + { + { + have a: A sorry + } + { + have b: B + proof - + { + show B sorry + } + { + have c: C sorry + } + { + have d: D sorry + } + qed + } + } + +end + + +section {* Calculational reasoning \label{sec:calculations-synopsis} *} + +text {* + For example, see @{file "~~/src/HOL/Isar_Examples/Group.thy"}. +*} + + +subsection {* Special names in Isar proofs *} + +text {* + \begin{itemize} + + \item term @{text "?thesis"} --- the main conclusion of the + innermost pending claim + + \item term @{text "\"} --- the argument of the last explicitly + stated result (for infix application this is the right-hand side) + + \item fact @{text "this"} --- the last result produced in the text + + \end{itemize} +*} + +notepad +begin + have "x = y" + proof - + term ?thesis + show ?thesis sorry + term ?thesis -- {* static! *} + qed + term "\" + thm this +end + +text {* Calculational reasoning maintains the special fact called + ``@{text calculation}'' in the background. Certain language + elements combine primary @{text this} with secondary @{text + calculation}. *} + + +subsection {* Transitive chains *} + +text {* The Idea is to combine @{text this} and @{text calculation} + via typical @{text trans} rules (see also @{command + print_trans_rules}): *} + +thm trans +thm less_trans +thm less_le_trans + +notepad +begin + txt {* Plain bottom-up calculation: *} + have "a = b" sorry + also + have "b = c" sorry + also + have "c = d" sorry + finally + have "a = d" . + + txt {* Variant using the @{text "\"} abbreviation: *} + have "a = b" sorry + also + have "\ = c" sorry + also + have "\ = d" sorry + finally + have "a = d" . + + txt {* Top-down version with explicit claim at the head: *} + have "a = d" + proof - + have "a = b" sorry + also + have "\ = c" sorry + also + have "\ = d" sorry + finally + show ?thesis . + qed +next + txt {* Mixed inequalities (require suitable base type): *} + fix a b c d :: nat + + have "a < b" sorry + also + have "b \ c" sorry + also + have "c = d" sorry + finally + have "a < d" . +end + + +subsubsection {* Notes *} + +text {* + \begin{itemize} + + \item The notion of @{text trans} rule is very general due to the + flexibility of Isabelle/Pure rule composition. + + \item User applications may declare their own rules, with some care + about the operational details of higher-order unification. + + \end{itemize} +*} + + +subsection {* Degenerate calculations and bigstep reasoning *} + +text {* The Idea is to append @{text this} to @{text calculation}, + without rule composition. *} + +notepad +begin + txt {* A vacuous proof: *} + have A sorry + moreover + have B sorry + moreover + have C sorry + ultimately + have A and B and C . +next + txt {* Slightly more content (trivial bigstep reasoning): *} + have A sorry + moreover + have B sorry + moreover + have C sorry + ultimately + have "A \ B \ C" by blast +next + txt {* More ambitious bigstep reasoning involving structured results: *} + have "A \ B \ C" sorry + moreover + { assume A have R sorry } + moreover + { assume B have R sorry } + moreover + { assume C have R sorry } + ultimately + have R by blast -- {* ``big-bang integration'' of proof blocks (occasionally fragile) *} +end + + +section {* Induction *} + +subsection {* Induction as Natural Deduction *} + +text {* In principle, induction is just a special case of Natural + Deduction (see also \secref{sec:natural-deduction-synopsis}). For + example: *} + +thm nat.induct +print_statement nat.induct + +notepad +begin + fix n :: nat + have "P n" + proof (rule nat.induct) -- {* fragile rule application! *} + show "P 0" sorry + next + fix n :: nat + assume "P n" + show "P (Suc n)" sorry + qed +end + +text {* + In practice, much more proof infrastructure is required. + + The proof method @{method induct} provides: + \begin{itemize} + + \item implicit rule selection and robust instantiation + + \item context elements via symbolic case names + + \item support for rule-structured induction statements, with local + parameters, premises, etc. + + \end{itemize} +*} + +notepad +begin + fix n :: nat + have "P n" + proof (induct n) + case 0 + show ?case sorry + next + case (Suc n) + from Suc.hyps show ?case sorry + qed +end + + +subsubsection {* Example *} + +text {* + The subsequent example combines the following proof patterns: + \begin{itemize} + + \item outermost induction (over the datatype structure of natural + numbers), to decompose the proof problem in top-down manner + + \item calculational reasoning (\secref{sec:calculations-synopsis}) + to compose the result in each case + + \item solving local claims within the calculation by simplification + + \end{itemize} +*} + +lemma + fixes n :: nat + shows "(\i=0..n. i) = n * (n + 1) div 2" +proof (induct n) + case 0 + have "(\i=0..0. i) = (0::nat)" by simp + also have "\ = 0 * (0 + 1) div 2" by simp + finally show ?case . +next + case (Suc n) + have "(\i=0..Suc n. i) = (\i=0..n. i) + (n + 1)" by simp + also have "\ = n * (n + 1) div 2 + (n + 1)" by (simp add: Suc.hyps) + also have "\ = (n * (n + 1) + 2 * (n + 1)) div 2" by simp + also have "\ = (Suc n * (Suc n + 1)) div 2" by simp + finally show ?case . +qed + +text {* This demonstrates how induction proofs can be done without + having to consider the raw Natural Deduction structure. *} + + +subsection {* Induction with local parameters and premises *} + +text {* Idea: Pure rule statements are passed through the induction + rule. This achieves convenient proof patterns, thanks to some + internal trickery in the @{method induct} method. + + Important: Using compact HOL formulae with @{text "\/\"} is a + well-known anti-pattern! It would produce useless formal noise. +*} + +notepad +begin + fix n :: nat + fix P :: "nat \ bool" + fix Q :: "'a \ nat \ bool" + + have "P n" + proof (induct n) + case 0 + show "P 0" sorry + next + case (Suc n) + from `P n` show "P (Suc n)" sorry + qed + + have "A n \ P n" + proof (induct n) + case 0 + from `A 0` show "P 0" sorry + next + case (Suc n) + from `A n \ P n` + and `A (Suc n)` show "P (Suc n)" sorry + qed + + have "\x. Q x n" + proof (induct n) + case 0 + show "Q x 0" sorry + next + case (Suc n) + from `\x. Q x n` show "Q x (Suc n)" sorry + txt {* Local quantification admits arbitrary instances: *} + note `Q a n` and `Q b n` + qed +end + + +subsection {* Implicit induction context *} + +text {* The @{method induct} method can isolate local parameters and + premises directly from the given statement. This is convenient in + practical applications, but requires some understanding of what is + going on internally (as explained above). *} + +notepad +begin + fix n :: nat + fix Q :: "'a \ nat \ bool" + + fix x :: 'a + assume "A x n" + then have "Q x n" + proof (induct n arbitrary: x) + case 0 + from `A x 0` show "Q x 0" sorry + next + case (Suc n) + from `\x. A x n \ Q x n` -- {* arbitrary instances can be produced here *} + and `A x (Suc n)` show "Q x (Suc n)" sorry + qed +end + + +subsection {* Advanced induction with term definitions *} + +text {* Induction over subexpressions of a certain shape are delicate + to formalize. The Isar @{method induct} method provides + infrastructure for this. + + Idea: sub-expressions of the problem are turned into a defined + induction variable; often accompanied with fixing of auxiliary + parameters in the original expression. *} + +notepad +begin + fix a :: "'a \ nat" + fix A :: "nat \ bool" + + assume "A (a x)" + then have "P (a x)" + proof (induct "a x" arbitrary: x) + case 0 + note prem = `A (a x)` + and defn = `0 = a x` + show "P (a x)" sorry + next + case (Suc n) + note hyp = `\x. n = a x \ A (a x) \ P (a x)` + and prem = `A (a x)` + and defn = `Suc n = a x` + show "P (a x)" sorry + qed +end + + +section {* Natural Deduction \label{sec:natural-deduction-synopsis} *} + +subsection {* Rule statements *} + +text {* + Isabelle/Pure ``theorems'' are always natural deduction rules, + which sometimes happen to consist of a conclusion only. + + The framework connectives @{text "\"} and @{text "\"} indicate the + rule structure declaratively. For example: *} + +thm conjI +thm impI +thm nat.induct + +text {* + The object-logic is embedded into the Pure framework via an implicit + derivability judgment @{term "Trueprop :: bool \ prop"}. + + Thus any HOL formulae appears atomic to the Pure framework, while + the rule structure outlines the corresponding proof pattern. + + This can be made explicit as follows: +*} + +notepad +begin + write Trueprop ("Tr") + + thm conjI + thm impI + thm nat.induct +end + +text {* + Isar provides first-class notation for rule statements as follows. +*} + +print_statement conjI +print_statement impI +print_statement nat.induct + + +subsubsection {* Examples *} + +text {* + Introductions and eliminations of some standard connectives of + the object-logic can be written as rule statements as follows. (The + proof ``@{command "by"}~@{method blast}'' serves as sanity check.) +*} + +lemma "(P \ False) \ \ P" by blast +lemma "\ P \ P \ Q" by blast + +lemma "P \ Q \ P \ Q" by blast +lemma "P \ Q \ (P \ Q \ R) \ R" by blast + +lemma "P \ P \ Q" by blast +lemma "Q \ P \ Q" by blast +lemma "P \ Q \ (P \ R) \ (Q \ R) \ R" by blast + +lemma "(\x. P x) \ (\x. P x)" by blast +lemma "(\x. P x) \ P x" by blast + +lemma "P x \ (\x. P x)" by blast +lemma "(\x. P x) \ (\x. P x \ R) \ R" by blast + +lemma "x \ A \ x \ B \ x \ A \ B" by blast +lemma "x \ A \ B \ (x \ A \ x \ B \ R) \ R" by blast + +lemma "x \ A \ x \ A \ B" by blast +lemma "x \ B \ x \ A \ B" by blast +lemma "x \ A \ B \ (x \ A \ R) \ (x \ B \ R) \ R" by blast + + +subsection {* Isar context elements *} + +text {* We derive some results out of the blue, using Isar context + elements and some explicit blocks. This illustrates their meaning + wrt.\ Pure connectives, without goal states getting in the way. *} + +notepad +begin + { + fix x + have "B x" sorry + } + have "\x. B x" by fact + +next + + { + assume A + have B sorry + } + have "A \ B" by fact + +next + + { + def x \ t + have "B x" sorry + } + have "B t" by fact + +next + + { + obtain x :: 'a where "B x" sorry + have C sorry + } + have C by fact + +end + + +subsection {* Pure rule composition *} + +text {* + The Pure framework provides means for: + + \begin{itemize} + + \item backward-chaining of rules by @{inference resolution} + + \item closing of branches by @{inference assumption} + + \end{itemize} + + Both principles involve higher-order unification of @{text \}-terms + modulo @{text "\\\"}-equivalence (cf.\ Huet and Miller). *} + +notepad +begin + assume a: A and b: B + thm conjI + thm conjI [of A B] -- "instantiation" + thm conjI [of A B, OF a b] -- "instantiation and composition" + thm conjI [OF a b] -- "composition via unification (trivial)" + thm conjI [OF `A` `B`] + + thm conjI [OF disjI1] +end + +text {* Note: Low-level rule composition is tedious and leads to + unreadable~/ unmaintainable expressions in the text. *} + + +subsection {* Structured backward reasoning *} + +text {* Idea: Canonical proof decomposition via @{command fix}~/ + @{command assume}~/ @{command show}, where the body produces a + natural deduction rule to refine some goal. *} + +notepad +begin + fix A B :: "'a \ bool" + + have "\x. A x \ B x" + proof - + fix x + assume "A x" + show "B x" sorry + qed + + have "\x. A x \ B x" + proof - + { + fix x + assume "A x" + show "B x" sorry + } -- "implicit block structure made explicit" + note `\x. A x \ B x` + -- "side exit for the resulting rule" + qed +end + + +subsection {* Structured rule application *} + +text {* + Idea: Previous facts and new claims are composed with a rule from + the context (or background library). +*} + +notepad +begin + assume r1: "A \ B \ C" -- {* simple rule (Horn clause) *} + + have A sorry -- "prefix of facts via outer sub-proof" + then have C + proof (rule r1) + show B sorry -- "remaining rule premises via inner sub-proof" + qed + + have C + proof (rule r1) + show A sorry + show B sorry + qed + + have A and B sorry + then have C + proof (rule r1) + qed + + have A and B sorry + then have C by (rule r1) + +next + + assume r2: "A \ (\x. B1 x \ B2 x) \ C" -- {* nested rule *} + + have A sorry + then have C + proof (rule r2) + fix x + assume "B1 x" + show "B2 x" sorry + qed + + txt {* The compound rule premise @{prop "\x. B1 x \ B2 x"} is better + addressed via @{command fix}~/ @{command assume}~/ @{command show} + in the nested proof body. *} +end + + +subsection {* Example: predicate logic *} + +text {* + Using the above principles, standard introduction and elimination proofs + of predicate logic connectives of HOL work as follows. +*} + +notepad +begin + have "A \ B" and A sorry + then have B .. + + have A sorry + then have "A \ B" .. + + have B sorry + then have "A \ B" .. + + have "A \ B" sorry + then have C + proof + assume A + then show C sorry + next + assume B + then show C sorry + qed + + have A and B sorry + then have "A \ B" .. + + have "A \ B" sorry + then have A .. + + have "A \ B" sorry + then have B .. + + have False sorry + then have A .. + + have True .. + + have "\ A" + proof + assume A + then show False sorry + qed + + have "\ A" and A sorry + then have B .. + + have "\x. P x" + proof + fix x + show "P x" sorry + qed + + have "\x. P x" sorry + then have "P a" .. + + have "\x. P x" + proof + show "P a" sorry + qed + + have "\x. P x" sorry + then have C + proof + fix a + assume "P a" + show C sorry + qed + + txt {* Less awkward version using @{command obtain}: *} + have "\x. P x" sorry + then obtain a where "P a" .. +end + +text {* Further variations to illustrate Isar sub-proofs involving + @{command show}: *} + +notepad +begin + have "A \ B" + proof -- {* two strictly isolated subproofs *} + show A sorry + next + show B sorry + qed + + have "A \ B" + proof -- {* one simultaneous sub-proof *} + show A and B sorry + qed + + have "A \ B" + proof -- {* two subproofs in the same context *} + show A sorry + show B sorry + qed + + have "A \ B" + proof -- {* swapped order *} + show B sorry + show A sorry + qed + + have "A \ B" + proof -- {* sequential subproofs *} + show A sorry + show B using `A` sorry + qed +end + + +subsubsection {* Example: set-theoretic operators *} + +text {* There is nothing special about logical connectives (@{text + "\"}, @{text "\"}, @{text "\"}, @{text "\"} etc.). Operators from + set-theory or lattice-theory work analogously. It is only a matter + of rule declarations in the library; rules can be also specified + explicitly. +*} + +notepad +begin + have "x \ A" and "x \ B" sorry + then have "x \ A \ B" .. + + have "x \ A" sorry + then have "x \ A \ B" .. + + have "x \ B" sorry + then have "x \ A \ B" .. + + have "x \ A \ B" sorry + then have C + proof + assume "x \ A" + then show C sorry + next + assume "x \ B" + then show C sorry + qed + +next + have "x \ \A" + proof + fix a + assume "a \ A" + show "x \ a" sorry + qed + + have "x \ \A" sorry + then have "x \ a" + proof + show "a \ A" sorry + qed + + have "a \ A" and "x \ a" sorry + then have "x \ \A" .. + + have "x \ \A" sorry + then obtain a where "a \ A" and "x \ a" .. +end + + +section {* Generalized elimination and cases *} + +subsection {* General elimination rules *} + +text {* + The general format of elimination rules is illustrated by the + following typical representatives: +*} + +thm exE -- {* local parameter *} +thm conjE -- {* local premises *} +thm disjE -- {* split into cases *} + +text {* + Combining these characteristics leads to the following general scheme + for elimination rules with cases: + + \begin{itemize} + + \item prefix of assumptions (or ``major premises'') + + \item one or more cases that enable to establish the main conclusion + in an augmented context + + \end{itemize} +*} + +notepad +begin + assume r: + "A1 \ A2 \ (* assumptions *) + (\x y. B1 x y \ C1 x y \ R) \ (* case 1 *) + (\x y. B2 x y \ C2 x y \ R) \ (* case 2 *) + R (* main conclusion *)" + + have A1 and A2 sorry + then have R + proof (rule r) + fix x y + assume "B1 x y" and "C1 x y" + show ?thesis sorry + next + fix x y + assume "B2 x y" and "C2 x y" + show ?thesis sorry + qed +end + +text {* Here @{text "?thesis"} is used to refer to the unchanged goal + statement. *} + + +subsection {* Rules with cases *} + +text {* + Applying an elimination rule to some goal, leaves that unchanged + but allows to augment the context in the sub-proof of each case. + + Isar provides some infrastructure to support this: + + \begin{itemize} + + \item native language elements to state eliminations + + \item symbolic case names + + \item method @{method cases} to recover this structure in a + sub-proof + + \end{itemize} +*} + +print_statement exE +print_statement conjE +print_statement disjE + +lemma + assumes A1 and A2 -- {* assumptions *} + obtains + (case1) x y where "B1 x y" and "C1 x y" + | (case2) x y where "B2 x y" and "C2 x y" + sorry + + +subsubsection {* Example *} + +lemma tertium_non_datur: + obtains + (T) A + | (F) "\ A" + by blast + +notepad +begin + fix x y :: 'a + have C + proof (cases "x = y" rule: tertium_non_datur) + case T + from `x = y` show ?thesis sorry + next + case F + from `x \ y` show ?thesis sorry + qed +end + + +subsubsection {* Example *} + +text {* + Isabelle/HOL specification mechanisms (datatype, inductive, etc.) + provide suitable derived cases rules. +*} + +datatype foo = Foo | Bar foo + +notepad +begin + fix x :: foo + have C + proof (cases x) + case Foo + from `x = Foo` show ?thesis sorry + next + case (Bar a) + from `x = Bar a` show ?thesis sorry + qed +end + + +subsection {* Obtaining local contexts *} + +text {* A single ``case'' branch may be inlined into Isar proof text + via @{command obtain}. This proves @{prop "(\x. B x \ thesis) \ + thesis"} on the spot, and augments the context afterwards. *} + +notepad +begin + fix B :: "'a \ bool" + + obtain x where "B x" sorry + note `B x` + + txt {* Conclusions from this context may not mention @{term x} again! *} + { + obtain x where "B x" sorry + from `B x` have C sorry + } + note `C` +end + +end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Base.thy --- a/doc-src/IsarRef/Thy/Base.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,8 +0,0 @@ -theory Base -imports Pure -begin - -ML_file "../../antiquote_setup.ML" -setup Antiquote_Setup.setup - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Document_Preparation.thy --- a/doc-src/IsarRef/Thy/Document_Preparation.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,611 +0,0 @@ -theory Document_Preparation -imports Base Main -begin - -chapter {* Document preparation \label{ch:document-prep} *} - -text {* Isabelle/Isar provides a simple document preparation system - based on regular {PDF-\LaTeX} technology, with full support for - hyper-links and bookmarks. Thus the results are well suited for WWW - browsing and as printed copies. - - \medskip Isabelle generates {\LaTeX} output while running a - \emph{logic session} (see also \cite{isabelle-sys}). Getting - started with a working configuration for common situations is quite - easy by using the Isabelle @{verbatim mkdir} and @{verbatim make} - tools. First invoke -\begin{ttbox} - isabelle mkdir Foo -\end{ttbox} - to initialize a separate directory for session @{verbatim Foo} (it - is safe to experiment, since @{verbatim "isabelle mkdir"} never - overwrites existing files). Ensure that @{verbatim "Foo/ROOT.ML"} - holds ML commands to load all theories required for this session; - furthermore @{verbatim "Foo/document/root.tex"} should include any - special {\LaTeX} macro packages required for your document (the - default is usually sufficient as a start). - - The session is controlled by a separate @{verbatim IsaMakefile} - (with crude source dependencies by default). This file is located - one level up from the @{verbatim Foo} directory location. Now - invoke -\begin{ttbox} - isabelle make Foo -\end{ttbox} - to run the @{verbatim Foo} session, with browser information and - document preparation enabled. Unless any errors are reported by - Isabelle or {\LaTeX}, the output will appear inside the directory - defined by the @{verbatim ISABELLE_BROWSER_INFO} setting (as - reported by the batch job in verbose mode). - - \medskip You may also consider to tune the @{verbatim usedir} - options in @{verbatim IsaMakefile}, for example to switch the output - format between @{verbatim pdf} and @{verbatim dvi}, or activate the - @{verbatim "-D"} option to retain a second copy of the generated - {\LaTeX} sources (for manual inspection or separate runs of - @{executable latex}). - - \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys} - for further details on Isabelle logic sessions and theory - presentation. The Isabelle/HOL tutorial \cite{isabelle-hol-book} - also covers theory presentation to some extent. -*} - - -section {* Markup commands \label{sec:markup} *} - -text {* - \begin{matharray}{rcl} - @{command_def "header"} & : & @{text "toplevel \ toplevel"} \\[0.5ex] - @{command_def "chapter"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "section"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "subsection"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "subsubsection"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "text"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "text_raw"} & : & @{text "local_theory \ local_theory"} \\[0.5ex] - @{command_def "sect"} & : & @{text "proof \ proof"} \\ - @{command_def "subsect"} & : & @{text "proof \ proof"} \\ - @{command_def "subsubsect"} & : & @{text "proof \ proof"} \\ - @{command_def "txt"} & : & @{text "proof \ proof"} \\ - @{command_def "txt_raw"} & : & @{text "proof \ proof"} \\ - \end{matharray} - - Markup commands provide a structured way to insert text into the - document generated from a theory. Each markup command takes a - single @{syntax text} argument, which is passed as argument to a - corresponding {\LaTeX} macro. The default macros provided by - @{file "~~/lib/texinputs/isabelle.sty"} can be redefined according - to the needs of the underlying document and {\LaTeX} styles. - - Note that formal comments (\secref{sec:comments}) are similar to - markup commands, but have a different status within Isabelle/Isar - syntax. - - @{rail " - (@@{command chapter} | @@{command section} | @@{command subsection} | - @@{command subsubsection} | @@{command text}) @{syntax target}? @{syntax text} - ; - (@@{command header} | @@{command text_raw} | @@{command sect} | @@{command subsect} | - @@{command subsubsect} | @@{command txt} | @@{command txt_raw}) @{syntax text} - "} - - \begin{description} - - \item @{command header} provides plain text markup just preceding - the formal beginning of a theory. The corresponding {\LaTeX} macro - is @{verbatim "\\isamarkupheader"}, which acts like @{command - section} by default. - - \item @{command chapter}, @{command section}, @{command subsection}, - and @{command subsubsection} mark chapter and section headings - within the main theory body or local theory targets. The - corresponding {\LaTeX} macros are @{verbatim "\\isamarkupchapter"}, - @{verbatim "\\isamarkupsection"}, @{verbatim - "\\isamarkupsubsection"} etc. - - \item @{command sect}, @{command subsect}, and @{command subsubsect} - mark section headings within proofs. The corresponding {\LaTeX} - macros are @{verbatim "\\isamarkupsect"}, @{verbatim - "\\isamarkupsubsect"} etc. - - \item @{command text} and @{command txt} specify paragraphs of plain - text. This corresponds to a {\LaTeX} environment @{verbatim - "\\begin{isamarkuptext}"} @{text "\"} @{verbatim - "\\end{isamarkuptext}"} etc. - - \item @{command text_raw} and @{command txt_raw} insert {\LaTeX} - source into the output, without additional markup. Thus the full - range of document manipulations becomes available, at the risk of - messing up document output. - - \end{description} - - Except for @{command "text_raw"} and @{command "txt_raw"}, the text - passed to any of the above markup commands may refer to formal - entities via \emph{document antiquotations}, see also - \secref{sec:antiq}. These are interpreted in the present theory or - proof context, or the named @{text "target"}. - - \medskip The proof markup commands closely resemble those for theory - specifications, but have a different formal status and produce - different {\LaTeX} macros. The default definitions coincide for - analogous commands such as @{command section} and @{command sect}. -*} - - -section {* Document Antiquotations \label{sec:antiq} *} - -text {* - \begin{matharray}{rcl} - @{antiquotation_def "theory"} & : & @{text antiquotation} \\ - @{antiquotation_def "thm"} & : & @{text antiquotation} \\ - @{antiquotation_def "lemma"} & : & @{text antiquotation} \\ - @{antiquotation_def "prop"} & : & @{text antiquotation} \\ - @{antiquotation_def "term"} & : & @{text antiquotation} \\ - @{antiquotation_def term_type} & : & @{text antiquotation} \\ - @{antiquotation_def typeof} & : & @{text antiquotation} \\ - @{antiquotation_def const} & : & @{text antiquotation} \\ - @{antiquotation_def abbrev} & : & @{text antiquotation} \\ - @{antiquotation_def typ} & : & @{text antiquotation} \\ - @{antiquotation_def type} & : & @{text antiquotation} \\ - @{antiquotation_def class} & : & @{text antiquotation} \\ - @{antiquotation_def "text"} & : & @{text antiquotation} \\ - @{antiquotation_def goals} & : & @{text antiquotation} \\ - @{antiquotation_def subgoals} & : & @{text antiquotation} \\ - @{antiquotation_def prf} & : & @{text antiquotation} \\ - @{antiquotation_def full_prf} & : & @{text antiquotation} \\ - @{antiquotation_def ML} & : & @{text antiquotation} \\ - @{antiquotation_def ML_op} & : & @{text antiquotation} \\ - @{antiquotation_def ML_type} & : & @{text antiquotation} \\ - @{antiquotation_def ML_struct} & : & @{text antiquotation} \\ - @{antiquotation_def "file"} & : & @{text antiquotation} \\ - \end{matharray} - - The overall content of an Isabelle/Isar theory may alternate between - formal and informal text. The main body consists of formal - specification and proof commands, interspersed with markup commands - (\secref{sec:markup}) or document comments (\secref{sec:comments}). - The argument of markup commands quotes informal text to be printed - in the resulting document, but may again refer to formal entities - via \emph{document antiquotations}. - - For example, embedding of ``@{text [source=false] "@{term [show_types] \"f x = a + x\"}"}'' - within a text block makes - \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document. - - Antiquotations usually spare the author tedious typing of logical - entities in full detail. Even more importantly, some degree of - consistency-checking between the main body of formal text and its - informal explanation is achieved, since terms and types appearing in - antiquotations are checked within the current theory or proof - context. - - %% FIXME less monolithic presentation, move to individual sections!? - @{rail " - '@{' antiquotation '}' - ; - @{syntax_def antiquotation}: - @@{antiquotation theory} options @{syntax name} | - @@{antiquotation thm} options styles @{syntax thmrefs} | - @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? | - @@{antiquotation prop} options styles @{syntax prop} | - @@{antiquotation term} options styles @{syntax term} | - @@{antiquotation (HOL) value} options styles @{syntax term} | - @@{antiquotation term_type} options styles @{syntax term} | - @@{antiquotation typeof} options styles @{syntax term} | - @@{antiquotation const} options @{syntax term} | - @@{antiquotation abbrev} options @{syntax term} | - @@{antiquotation typ} options @{syntax type} | - @@{antiquotation type} options @{syntax name} | - @@{antiquotation class} options @{syntax name} | - @@{antiquotation text} options @{syntax name} - ; - @{syntax antiquotation}: - @@{antiquotation goals} options | - @@{antiquotation subgoals} options | - @@{antiquotation prf} options @{syntax thmrefs} | - @@{antiquotation full_prf} options @{syntax thmrefs} | - @@{antiquotation ML} options @{syntax name} | - @@{antiquotation ML_op} options @{syntax name} | - @@{antiquotation ML_type} options @{syntax name} | - @@{antiquotation ML_struct} options @{syntax name} | - @@{antiquotation \"file\"} options @{syntax name} - ; - options: '[' (option * ',') ']' - ; - option: @{syntax name} | @{syntax name} '=' @{syntax name} - ; - styles: '(' (style + ',') ')' - ; - style: (@{syntax name} +) - "} - - Note that the syntax of antiquotations may \emph{not} include source - comments @{verbatim "(*"}~@{text "\"}~@{verbatim "*)"} nor verbatim - text @{verbatim "{"}@{verbatim "*"}~@{text "\"}~@{verbatim - "*"}@{verbatim "}"}. - - \begin{description} - - \item @{text "@{theory A}"} prints the name @{text "A"}, which is - guaranteed to refer to a valid ancestor theory in the current - context. - - \item @{text "@{thm a\<^sub>1 \ a\<^sub>n}"} prints theorems @{text "a\<^sub>1 \ a\<^sub>n"}. - Full fact expressions are allowed here, including attributes - (\secref{sec:syn-att}). - - \item @{text "@{prop \}"} prints a well-typed proposition @{text - "\"}. - - \item @{text "@{lemma \ by m}"} proves a well-typed proposition - @{text "\"} by method @{text m} and prints the original @{text "\"}. - - \item @{text "@{term t}"} prints a well-typed term @{text "t"}. - - \item @{text "@{value t}"} evaluates a term @{text "t"} and prints - its result, see also @{command_ref (HOL) value}. - - \item @{text "@{term_type t}"} prints a well-typed term @{text "t"} - annotated with its type. - - \item @{text "@{typeof t}"} prints the type of a well-typed term - @{text "t"}. - - \item @{text "@{const c}"} prints a logical or syntactic constant - @{text "c"}. - - \item @{text "@{abbrev c x\<^sub>1 \ x\<^sub>n}"} prints a constant abbreviation - @{text "c x\<^sub>1 \ x\<^sub>n \ rhs"} as defined in the current context. - - \item @{text "@{typ \}"} prints a well-formed type @{text "\"}. - - \item @{text "@{type \}"} prints a (logical or syntactic) type - constructor @{text "\"}. - - \item @{text "@{class c}"} prints a class @{text c}. - - \item @{text "@{text s}"} prints uninterpreted source text @{text - s}. This is particularly useful to print portions of text according - to the Isabelle document style, without demanding well-formedness, - e.g.\ small pieces of terms that should not be parsed or - type-checked yet. - - \item @{text "@{goals}"} prints the current \emph{dynamic} goal - state. This is mainly for support of tactic-emulation scripts - within Isar. Presentation of goal states does not conform to the - idea of human-readable proof documents! - - When explaining proofs in detail it is usually better to spell out - the reasoning via proper Isar proof commands, instead of peeking at - the internal machine configuration. - - \item @{text "@{subgoals}"} is similar to @{text "@{goals}"}, but - does not print the main goal. - - \item @{text "@{prf a\<^sub>1 \ a\<^sub>n}"} prints the (compact) proof terms - corresponding to the theorems @{text "a\<^sub>1 \ a\<^sub>n"}. Note that this - requires proof terms to be switched on for the current logic - session. - - \item @{text "@{full_prf a\<^sub>1 \ a\<^sub>n}"} is like @{text "@{prf a\<^sub>1 \ - a\<^sub>n}"}, but prints the full proof terms, i.e.\ also displays - information omitted in the compact proof term, which is denoted by - ``@{text _}'' placeholders there. - - \item @{text "@{ML s}"}, @{text "@{ML_op s}"}, @{text "@{ML_type - s}"}, and @{text "@{ML_struct s}"} check text @{text s} as ML value, - infix operator, type, and structure, respectively. The source is - printed verbatim. - - \item @{text "@{file path}"} checks that @{text "path"} refers to a - file (or directory) and prints it verbatim. - - \end{description} -*} - - -subsection {* Styled antiquotations *} - -text {* The antiquotations @{text thm}, @{text prop} and @{text - term} admit an extra \emph{style} specification to modify the - printed result. A style is specified by a name with a possibly - empty number of arguments; multiple styles can be sequenced with - commas. The following standard styles are available: - - \begin{description} - - \item @{text lhs} extracts the first argument of any application - form with at least two arguments --- typically meta-level or - object-level equality, or any other binary relation. - - \item @{text rhs} is like @{text lhs}, but extracts the second - argument. - - \item @{text "concl"} extracts the conclusion @{text C} from a rule - in Horn-clause normal form @{text "A\<^sub>1 \ \ A\<^sub>n \ C"}. - - \item @{text "prem"} @{text n} extract premise number - @{text "n"} from from a rule in Horn-clause - normal form @{text "A\<^sub>1 \ \ A\<^sub>n \ C"} - - \end{description} -*} - - -subsection {* General options *} - -text {* The following options are available to tune the printed output - of antiquotations. Note that many of these coincide with global ML - flags of the same names. - - \begin{description} - - \item @{antiquotation_option_def show_types}~@{text "= bool"} and - @{antiquotation_option_def show_sorts}~@{text "= bool"} control - printing of explicit type and sort constraints. - - \item @{antiquotation_option_def show_structs}~@{text "= bool"} - controls printing of implicit structures. - - \item @{antiquotation_option_def show_abbrevs}~@{text "= bool"} - controls folding of abbreviations. - - \item @{antiquotation_option_def names_long}~@{text "= bool"} forces - names of types and constants etc.\ to be printed in their fully - qualified internal form. - - \item @{antiquotation_option_def names_short}~@{text "= bool"} - forces names of types and constants etc.\ to be printed unqualified. - Note that internalizing the output again in the current context may - well yield a different result. - - \item @{antiquotation_option_def names_unique}~@{text "= bool"} - determines whether the printed version of qualified names should be - made sufficiently long to avoid overlap with names declared further - back. Set to @{text false} for more concise output. - - \item @{antiquotation_option_def eta_contract}~@{text "= bool"} - prints terms in @{text \}-contracted form. - - \item @{antiquotation_option_def display}~@{text "= bool"} indicates - if the text is to be output as multi-line ``display material'', - rather than a small piece of text without line breaks (which is the - default). - - In this mode the embedded entities are printed in the same style as - the main theory text. - - \item @{antiquotation_option_def break}~@{text "= bool"} controls - line breaks in non-display material. - - \item @{antiquotation_option_def quotes}~@{text "= bool"} indicates - if the output should be enclosed in double quotes. - - \item @{antiquotation_option_def mode}~@{text "= name"} adds @{text - name} to the print mode to be used for presentation. Note that the - standard setup for {\LaTeX} output is already present by default, - including the modes @{text latex} and @{text xsymbols}. - - \item @{antiquotation_option_def margin}~@{text "= nat"} and - @{antiquotation_option_def indent}~@{text "= nat"} change the margin - or indentation for pretty printing of display material. - - \item @{antiquotation_option_def goals_limit}~@{text "= nat"} - determines the maximum number of goals to be printed (for goal-based - antiquotation). - - \item @{antiquotation_option_def source}~@{text "= bool"} prints the - original source text of the antiquotation arguments, rather than its - internal representation. Note that formal checking of - @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still - enabled; use the @{antiquotation "text"} antiquotation for unchecked - output. - - Regular @{text "term"} and @{text "typ"} antiquotations with @{text - "source = false"} involve a full round-trip from the original source - to an internalized logical entity back to a source form, according - to the syntax of the current context. Thus the printed output is - not under direct control of the author, it may even fluctuate a bit - as the underlying theory is changed later on. - - In contrast, @{antiquotation_option source}~@{text "= true"} - admits direct printing of the given source text, with the desirable - well-formedness check in the background, but without modification of - the printed text. - - \end{description} - - For boolean flags, ``@{text "name = true"}'' may be abbreviated as - ``@{text name}''. All of the above flags are disabled by default, - unless changed from ML, say in the @{verbatim "ROOT.ML"} of the - logic session. -*} - - -section {* Markup via command tags \label{sec:tags} *} - -text {* Each Isabelle/Isar command may be decorated by additional - presentation tags, to indicate some modification in the way it is - printed in the document. - - @{rail " - @{syntax_def tags}: ( tag * ) - ; - tag: '%' (@{syntax ident} | @{syntax string}) - "} - - Some tags are pre-declared for certain classes of commands, serving - as default markup if no tags are given in the text: - - \medskip - \begin{tabular}{ll} - @{text "theory"} & theory begin/end \\ - @{text "proof"} & all proof commands \\ - @{text "ML"} & all commands involving ML code \\ - \end{tabular} - - \medskip The Isabelle document preparation system - \cite{isabelle-sys} allows tagged command regions to be presented - specifically, e.g.\ to fold proof texts, or drop parts of the text - completely. - - For example ``@{command "by"}~@{text "%invisible auto"}'' causes - that piece of proof to be treated as @{text invisible} instead of - @{text "proof"} (the default), which may be shown or hidden - depending on the document setup. In contrast, ``@{command - "by"}~@{text "%visible auto"}'' forces this text to be shown - invariably. - - Explicit tag specifications within a proof apply to all subsequent - commands of the same level of nesting. For example, ``@{command - "proof"}~@{text "%visible \"}~@{command "qed"}'' forces the whole - sub-proof to be typeset as @{text visible} (unless some of its parts - are tagged differently). - - \medskip Command tags merely produce certain markup environments for - type-setting. The meaning of these is determined by {\LaTeX} - macros, as defined in @{file "~~/lib/texinputs/isabelle.sty"} or - by the document author. The Isabelle document preparation tools - also provide some high-level options to specify the meaning of - arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding - parts of the text. Logic sessions may also specify ``document - versions'', where given tags are interpreted in some particular way. - Again see \cite{isabelle-sys} for further details. -*} - - -section {* Railroad diagrams *} - -text {* - \begin{matharray}{rcl} - @{antiquotation_def "rail"} & : & @{text antiquotation} \\ - \end{matharray} - - @{rail "'rail' string"} - - The @{antiquotation rail} antiquotation allows to include syntax - diagrams into Isabelle documents. {\LaTeX} requires the style file - @{"file" "~~/lib/texinputs/pdfsetup.sty"}, which can be used via - @{verbatim "\\usepackage{pdfsetup}"} in @{verbatim "root.tex"}, for - example. - - The rail specification language is quoted here as Isabelle @{syntax - string}; it has its own grammar given below. - - @{rail " - rule? + ';' - ; - rule: ((identifier | @{syntax antiquotation}) ':')? body - ; - body: concatenation + '|' - ; - concatenation: ((atom '?'?) +) (('*' | '+') atom?)? - ; - atom: '(' body? ')' | identifier | - '@'? (string | @{syntax antiquotation}) | - '\\\\\\\\' - "} - - The lexical syntax of @{text "identifier"} coincides with that of - @{syntax ident} in regular Isabelle syntax, but @{text string} uses - single quotes instead of double quotes of the standard @{syntax - string} category, to avoid extra escapes. - - Each @{text rule} defines a formal language (with optional name), - using a notation that is similar to EBNF or regular expressions with - recursion. The meaning and visual appearance of these rail language - elements is illustrated by the following representative examples. - - \begin{itemize} - - \item Empty @{verbatim "()"} - - @{rail "()"} - - \item Nonterminal @{verbatim "A"} - - @{rail "A"} - - \item Nonterminal via Isabelle antiquotation - @{verbatim "@{syntax method}"} - - @{rail "@{syntax method}"} - - \item Terminal @{verbatim "'xyz'"} - - @{rail "'xyz'"} - - \item Terminal in keyword style @{verbatim "@'xyz'"} - - @{rail "@'xyz'"} - - \item Terminal via Isabelle antiquotation - @{verbatim "@@{method rule}"} - - @{rail "@@{method rule}"} - - \item Concatenation @{verbatim "A B C"} - - @{rail "A B C"} - - \item Linebreak @{verbatim "\\\\"} inside - concatenation\footnote{Strictly speaking, this is only a single - backslash, but the enclosing @{syntax string} syntax requires a - second one for escaping.} @{verbatim "A B C \\\\ D E F"} - - @{rail "A B C \\ D E F"} - - \item Variants @{verbatim "A | B | C"} - - @{rail "A | B | C"} - - \item Option @{verbatim "A ?"} - - @{rail "A ?"} - - \item Repetition @{verbatim "A *"} - - @{rail "A *"} - - \item Repetition with separator @{verbatim "A * sep"} - - @{rail "A * sep"} - - \item Strict repetition @{verbatim "A +"} - - @{rail "A +"} - - \item Strict repetition with separator @{verbatim "A + sep"} - - @{rail "A + sep"} - - \end{itemize} -*} - - -section {* Draft presentation *} - -text {* - \begin{matharray}{rcl} - @{command_def "display_drafts"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - @{command_def "print_drafts"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - \end{matharray} - - @{rail " - (@@{command display_drafts} | @@{command print_drafts}) (@{syntax name} +) - - "} - - \begin{description} - - \item @{command "display_drafts"}~@{text paths} and @{command - "print_drafts"}~@{text paths} perform simple output of a given list - of raw source files. Only those symbols that do not require - additional {\LaTeX} packages are displayed properly, everything else - is left verbatim. - - \end{description} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/First_Order_Logic.thy --- a/doc-src/IsarRef/Thy/First_Order_Logic.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,520 +0,0 @@ - -header {* Example: First-Order Logic *} - -theory %visible First_Order_Logic -imports Base (* FIXME Pure!? *) -begin - -text {* - \noindent In order to commence a new object-logic within - Isabelle/Pure we introduce abstract syntactic categories @{text "i"} - for individuals and @{text "o"} for object-propositions. The latter - is embedded into the language of Pure propositions by means of a - separate judgment. -*} - -typedecl i -typedecl o - -judgment - Trueprop :: "o \ prop" ("_" 5) - -text {* - \noindent Note that the object-logic judgement is implicit in the - syntax: writing @{prop A} produces @{term "Trueprop A"} internally. - From the Pure perspective this means ``@{prop A} is derivable in the - object-logic''. -*} - - -subsection {* Equational reasoning \label{sec:framework-ex-equal} *} - -text {* - Equality is axiomatized as a binary predicate on individuals, with - reflexivity as introduction, and substitution as elimination - principle. Note that the latter is particularly convenient in a - framework like Isabelle, because syntactic congruences are - implicitly produced by unification of @{term "B x"} against - expressions containing occurrences of @{term x}. -*} - -axiomatization - equal :: "i \ i \ o" (infix "=" 50) -where - refl [intro]: "x = x" and - subst [elim]: "x = y \ B x \ B y" - -text {* - \noindent Substitution is very powerful, but also hard to control in - full generality. We derive some common symmetry~/ transitivity - schemes of @{term equal} as particular consequences. -*} - -theorem sym [sym]: - assumes "x = y" - shows "y = x" -proof - - have "x = x" .. - with `x = y` show "y = x" .. -qed - -theorem forw_subst [trans]: - assumes "y = x" and "B x" - shows "B y" -proof - - from `y = x` have "x = y" .. - from this and `B x` show "B y" .. -qed - -theorem back_subst [trans]: - assumes "B x" and "x = y" - shows "B y" -proof - - from `x = y` and `B x` - show "B y" .. -qed - -theorem trans [trans]: - assumes "x = y" and "y = z" - shows "x = z" -proof - - from `y = z` and `x = y` - show "x = z" .. -qed - - -subsection {* Basic group theory *} - -text {* - As an example for equational reasoning we consider some bits of - group theory. The subsequent locale definition postulates group - operations and axioms; we also derive some consequences of this - specification. -*} - -locale group = - fixes prod :: "i \ i \ i" (infix "\" 70) - and inv :: "i \ i" ("(_\)" [1000] 999) - and unit :: i ("1") - assumes assoc: "(x \ y) \ z = x \ (y \ z)" - and left_unit: "1 \ x = x" - and left_inv: "x\ \ x = 1" -begin - -theorem right_inv: "x \ x\ = 1" -proof - - have "x \ x\ = 1 \ (x \ x\)" by (rule left_unit [symmetric]) - also have "\ = (1 \ x) \ x\" by (rule assoc [symmetric]) - also have "1 = (x\)\ \ x\" by (rule left_inv [symmetric]) - also have "\ \ x = (x\)\ \ (x\ \ x)" by (rule assoc) - also have "x\ \ x = 1" by (rule left_inv) - also have "((x\)\ \ \) \ x\ = (x\)\ \ (1 \ x\)" by (rule assoc) - also have "1 \ x\ = x\" by (rule left_unit) - also have "(x\)\ \ \ = 1" by (rule left_inv) - finally show "x \ x\ = 1" . -qed - -theorem right_unit: "x \ 1 = x" -proof - - have "1 = x\ \ x" by (rule left_inv [symmetric]) - also have "x \ \ = (x \ x\) \ x" by (rule assoc [symmetric]) - also have "x \ x\ = 1" by (rule right_inv) - also have "\ \ x = x" by (rule left_unit) - finally show "x \ 1 = x" . -qed - -text {* - \noindent Reasoning from basic axioms is often tedious. Our proofs - work by producing various instances of the given rules (potentially - the symmetric form) using the pattern ``@{command have}~@{text - eq}~@{command "by"}~@{text "(rule r)"}'' and composing the chain of - results via @{command also}/@{command finally}. These steps may - involve any of the transitivity rules declared in - \secref{sec:framework-ex-equal}, namely @{thm trans} in combining - the first two results in @{thm right_inv} and in the final steps of - both proofs, @{thm forw_subst} in the first combination of @{thm - right_unit}, and @{thm back_subst} in all other calculational steps. - - Occasional substitutions in calculations are adequate, but should - not be over-emphasized. The other extreme is to compose a chain by - plain transitivity only, with replacements occurring always in - topmost position. For example: -*} - -(*<*) -theorem "\A. PROP A \ PROP A" -proof - - assume [symmetric, defn]: "\x y. (x \ y) \ Trueprop (x = y)" -(*>*) - have "x \ 1 = x \ (x\ \ x)" unfolding left_inv .. - also have "\ = (x \ x\) \ x" unfolding assoc .. - also have "\ = 1 \ x" unfolding right_inv .. - also have "\ = x" unfolding left_unit .. - finally have "x \ 1 = x" . -(*<*) -qed -(*>*) - -text {* - \noindent Here we have re-used the built-in mechanism for unfolding - definitions in order to normalize each equational problem. A more - realistic object-logic would include proper setup for the Simplifier - (\secref{sec:simplifier}), the main automated tool for equational - reasoning in Isabelle. Then ``@{command unfolding}~@{thm - left_inv}~@{command ".."}'' would become ``@{command "by"}~@{text - "(simp only: left_inv)"}'' etc. -*} - -end - - -subsection {* Propositional logic \label{sec:framework-ex-prop} *} - -text {* - We axiomatize basic connectives of propositional logic: implication, - disjunction, and conjunction. The associated rules are modeled - after Gentzen's system of Natural Deduction \cite{Gentzen:1935}. -*} - -axiomatization - imp :: "o \ o \ o" (infixr "\" 25) where - impI [intro]: "(A \ B) \ A \ B" and - impD [dest]: "(A \ B) \ A \ B" - -axiomatization - disj :: "o \ o \ o" (infixr "\" 30) where - disjI\<^isub>1 [intro]: "A \ A \ B" and - disjI\<^isub>2 [intro]: "B \ A \ B" and - disjE [elim]: "A \ B \ (A \ C) \ (B \ C) \ C" - -axiomatization - conj :: "o \ o \ o" (infixr "\" 35) where - conjI [intro]: "A \ B \ A \ B" and - conjD\<^isub>1: "A \ B \ A" and - conjD\<^isub>2: "A \ B \ B" - -text {* - \noindent The conjunctive destructions have the disadvantage that - decomposing @{prop "A \ B"} involves an immediate decision which - component should be projected. The more convenient simultaneous - elimination @{prop "A \ B \ (A \ B \ C) \ C"} can be derived as - follows: -*} - -theorem conjE [elim]: - assumes "A \ B" - obtains A and B -proof - from `A \ B` show A by (rule conjD\<^isub>1) - from `A \ B` show B by (rule conjD\<^isub>2) -qed - -text {* - \noindent Here is an example of swapping conjuncts with a single - intermediate elimination step: -*} - -(*<*) -lemma "\A. PROP A \ PROP A" -proof - -(*>*) - assume "A \ B" - then obtain B and A .. - then have "B \ A" .. -(*<*) -qed -(*>*) - -text {* - \noindent Note that the analogous elimination rule for disjunction - ``@{text "\ A \ B \ A \ B"}'' coincides with - the original axiomatization of @{thm disjE}. - - \medskip We continue propositional logic by introducing absurdity - with its characteristic elimination. Plain truth may then be - defined as a proposition that is trivially true. -*} - -axiomatization - false :: o ("\") where - falseE [elim]: "\ \ A" - -definition - true :: o ("\") where - "\ \ \ \ \" - -theorem trueI [intro]: \ - unfolding true_def .. - -text {* - \medskip\noindent Now negation represents an implication towards - absurdity: -*} - -definition - not :: "o \ o" ("\ _" [40] 40) where - "\ A \ A \ \" - -theorem notI [intro]: - assumes "A \ \" - shows "\ A" -unfolding not_def -proof - assume A - then show \ by (rule `A \ \`) -qed - -theorem notE [elim]: - assumes "\ A" and A - shows B -proof - - from `\ A` have "A \ \" unfolding not_def . - from `A \ \` and `A` have \ .. - then show B .. -qed - - -subsection {* Classical logic *} - -text {* - Subsequently we state the principle of classical contradiction as a - local assumption. Thus we refrain from forcing the object-logic - into the classical perspective. Within that context, we may derive - well-known consequences of the classical principle. -*} - -locale classical = - assumes classical: "(\ C \ C) \ C" -begin - -theorem double_negation: - assumes "\ \ C" - shows C -proof (rule classical) - assume "\ C" - with `\ \ C` show C .. -qed - -theorem tertium_non_datur: "C \ \ C" -proof (rule double_negation) - show "\ \ (C \ \ C)" - proof - assume "\ (C \ \ C)" - have "\ C" - proof - assume C then have "C \ \ C" .. - with `\ (C \ \ C)` show \ .. - qed - then have "C \ \ C" .. - with `\ (C \ \ C)` show \ .. - qed -qed - -text {* - \noindent These examples illustrate both classical reasoning and - non-trivial propositional proofs in general. All three rules - characterize classical logic independently, but the original rule is - already the most convenient to use, because it leaves the conclusion - unchanged. Note that @{prop "(\ C \ C) \ C"} fits again into our - format for eliminations, despite the additional twist that the - context refers to the main conclusion. So we may write @{thm - classical} as the Isar statement ``@{text "\ \ thesis"}''. - This also explains nicely how classical reasoning really works: - whatever the main @{text thesis} might be, we may always assume its - negation! -*} - -end - - -subsection {* Quantifiers \label{sec:framework-ex-quant} *} - -text {* - Representing quantifiers is easy, thanks to the higher-order nature - of the underlying framework. According to the well-known technique - introduced by Church \cite{church40}, quantifiers are operators on - predicates, which are syntactically represented as @{text "\"}-terms - of type @{typ "i \ o"}. Binder notation turns @{text "All (\x. B - x)"} into @{text "\x. B x"} etc. -*} - -axiomatization - All :: "(i \ o) \ o" (binder "\" 10) where - allI [intro]: "(\x. B x) \ \x. B x" and - allD [dest]: "(\x. B x) \ B a" - -axiomatization - Ex :: "(i \ o) \ o" (binder "\" 10) where - exI [intro]: "B a \ (\x. B x)" and - exE [elim]: "(\x. B x) \ (\x. B x \ C) \ C" - -text {* - \noindent The statement of @{thm exE} corresponds to ``@{text - "\ \x. B x \ x \ B x"}'' in Isar. In the - subsequent example we illustrate quantifier reasoning involving all - four rules: -*} - -theorem - assumes "\x. \y. R x y" - shows "\y. \x. R x y" -proof -- {* @{text "\"} introduction *} - obtain x where "\y. R x y" using `\x. \y. R x y` .. -- {* @{text "\"} elimination *} - fix y have "R x y" using `\y. R x y` .. -- {* @{text "\"} destruction *} - then show "\x. R x y" .. -- {* @{text "\"} introduction *} -qed - - -subsection {* Canonical reasoning patterns *} - -text {* - The main rules of first-order predicate logic from - \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} - can now be summarized as follows, using the native Isar statement - format of \secref{sec:framework-stmt}. - - \medskip - \begin{tabular}{l} - @{text "impI: \ A \ B \ A \ B"} \\ - @{text "impD: \ A \ B \ A \ B"} \\[1ex] - - @{text "disjI\<^isub>1: \ A \ A \ B"} \\ - @{text "disjI\<^isub>2: \ B \ A \ B"} \\ - @{text "disjE: \ A \ B \ A \ B"} \\[1ex] - - @{text "conjI: \ A \ B \ A \ B"} \\ - @{text "conjE: \ A \ B \ A \ B"} \\[1ex] - - @{text "falseE: \ \ \ A"} \\ - @{text "trueI: \ \"} \\[1ex] - - @{text "notI: \ A \ \ \ \ A"} \\ - @{text "notE: \ \ A \ A \ B"} \\[1ex] - - @{text "allI: \ \x. B x \ \x. B x"} \\ - @{text "allE: \ \x. B x \ B a"} \\[1ex] - - @{text "exI: \ B a \ \x. B x"} \\ - @{text "exE: \ \x. B x \ a \ B a"} - \end{tabular} - \medskip - - \noindent This essentially provides a declarative reading of Pure - rules as Isar reasoning patterns: the rule statements tells how a - canonical proof outline shall look like. Since the above rules have - already been declared as @{attribute (Pure) intro}, @{attribute - (Pure) elim}, @{attribute (Pure) dest} --- each according to its - particular shape --- we can immediately write Isar proof texts as - follows: -*} - -(*<*) -theorem "\A. PROP A \ PROP A" -proof - -(*>*) - - txt_raw {*\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "A \ B" - proof - assume A - show B sorry %noproof - qed - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "A \ B" and A sorry %noproof - then have B .. - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have A sorry %noproof - then have "A \ B" .. - - have B sorry %noproof - then have "A \ B" .. - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "A \ B" sorry %noproof - then have C - proof - assume A - then show C sorry %noproof - next - assume B - then show C sorry %noproof - qed - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have A and B sorry %noproof - then have "A \ B" .. - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "A \ B" sorry %noproof - then obtain A and B .. - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\" sorry %noproof - then have A .. - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\" .. - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\ A" - proof - assume A - then show "\" sorry %noproof - qed - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\ A" and A sorry %noproof - then have B .. - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\x. B x" - proof - fix x - show "B x" sorry %noproof - qed - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\x. B x" sorry %noproof - then have "B a" .. - - txt_raw {*\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\x. B x" - proof - show "B a" sorry %noproof - qed - - txt_raw {*\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}*}(*<*)next(*>*) - - have "\x. B x" sorry %noproof - then obtain a where "B a" .. - - txt_raw {*\end{minipage}*} - -(*<*) -qed -(*>*) - -text {* - \bigskip\noindent Of course, these proofs are merely examples. As - sketched in \secref{sec:framework-subproof}, there is a fair amount - of flexibility in expressing Pure deductions in Isar. Here the user - is asked to express himself adequately, aiming at proof texts of - literary quality. -*} - -end %visible diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Framework.thy --- a/doc-src/IsarRef/Thy/Framework.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1016 +0,0 @@ -theory Framework -imports Base Main -begin - -chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *} - -text {* - Isabelle/Isar - \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift} - is intended as a generic framework for developing formal - mathematical documents with full proof checking. Definitions and - proofs are organized as theories. An assembly of theory sources may - be presented as a printed document; see also - \chref{ch:document-prep}. - - The main objective of Isar is the design of a human-readable - structured proof language, which is called the ``primary proof - format'' in Isar terminology. Such a primary proof language is - somewhere in the middle between the extremes of primitive proof - objects and actual natural language. In this respect, Isar is a bit - more formalistic than Mizar - \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar}, - using logical symbols for certain reasoning schemes where Mizar - would prefer English words; see \cite{Wenzel-Wiedijk:2002} for - further comparisons of these systems. - - So Isar challenges the traditional way of recording informal proofs - in mathematical prose, as well as the common tendency to see fully - formal proofs directly as objects of some logical calculus (e.g.\ - @{text "\"}-terms in a version of type theory). In fact, Isar is - better understood as an interpreter of a simple block-structured - language for describing the data flow of local facts and goals, - interspersed with occasional invocations of proof methods. - Everything is reduced to logical inferences internally, but these - steps are somewhat marginal compared to the overall bookkeeping of - the interpretation process. Thanks to careful design of the syntax - and semantics of Isar language elements, a formal record of Isar - instructions may later appear as an intelligible text to the - attentive reader. - - The Isar proof language has emerged from careful analysis of some - inherent virtues of the existing logical framework of Isabelle/Pure - \cite{paulson-found,paulson700}, notably composition of higher-order - natural deduction rules, which is a generalization of Gentzen's - original calculus \cite{Gentzen:1935}. The approach of generic - inference systems in Pure is continued by Isar towards actual proof - texts. - - Concrete applications require another intermediate layer: an - object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed - set-theory) is being used most of the time; Isabelle/ZF - \cite{isabelle-ZF} is less extensively developed, although it would - probably fit better for classical mathematics. - - \medskip In order to illustrate natural deduction in Isar, we shall - refer to the background theory and library of Isabelle/HOL. This - includes common notions of predicate logic, naive set-theory etc.\ - using fairly standard mathematical notation. From the perspective - of generic natural deduction there is nothing special about the - logical connectives of HOL (@{text "\"}, @{text "\"}, @{text "\"}, - @{text "\"}, etc.), only the resulting reasoning principles are - relevant to the user. There are similar rules available for - set-theory operators (@{text "\"}, @{text "\"}, @{text "\"}, @{text - "\"}, etc.), or any other theory developed in the library (lattice - theory, topology etc.). - - Subsequently we briefly review fragments of Isar proof texts - corresponding directly to such general deduction schemes. The - examples shall refer to set-theory, to minimize the danger of - understanding connectives of predicate logic as something special. - - \medskip The following deduction performs @{text "\"}-introduction, - working forwards from assumptions towards the conclusion. We give - both the Isar text, and depict the primitive rule involved, as - determined by unification of the problem against rules that are - declared in the library context. -*} - -text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} - -(*<*) -notepad -begin -(*>*) - assume "x \ A" and "x \ B" - then have "x \ A \ B" .. -(*<*) -end -(*>*) - -text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} - -text {* - \infer{@{prop "x \ A \ B"}}{@{prop "x \ A"} & @{prop "x \ B"}} -*} - -text_raw {*\end{minipage}*} - -text {* - \medskip\noindent Note that @{command assume} augments the proof - context, @{command then} indicates that the current fact shall be - used in the next step, and @{command have} states an intermediate - goal. The two dots ``@{command ".."}'' refer to a complete proof of - this claim, using the indicated facts and a canonical rule from the - context. We could have been more explicit here by spelling out the - final proof step via the @{command "by"} command: -*} - -(*<*) -notepad -begin -(*>*) - assume "x \ A" and "x \ B" - then have "x \ A \ B" by (rule IntI) -(*<*) -end -(*>*) - -text {* - \noindent The format of the @{text "\"}-introduction rule represents - the most basic inference, which proceeds from given premises to a - conclusion, without any nested proof context involved. - - The next example performs backwards introduction on @{term "\\"}, - the intersection of all sets within a given set. This requires a - nested proof of set membership within a local context, where @{term - A} is an arbitrary-but-fixed member of the collection: -*} - -text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} - -(*<*) -notepad -begin -(*>*) - have "x \ \\" - proof - fix A - assume "A \ \" - show "x \ A" sorry %noproof - qed -(*<*) -end -(*>*) - -text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} - -text {* - \infer{@{prop "x \ \\"}}{\infer*{@{prop "x \ A"}}{@{text "[A][A \ \]"}}} -*} - -text_raw {*\end{minipage}*} - -text {* - \medskip\noindent This Isar reasoning pattern again refers to the - primitive rule depicted above. The system determines it in the - ``@{command proof}'' step, which could have been spelt out more - explicitly as ``@{command proof}~@{text "(rule InterI)"}''. Note - that the rule involves both a local parameter @{term "A"} and an - assumption @{prop "A \ \"} in the nested reasoning. This kind of - compound rule typically demands a genuine sub-proof in Isar, working - backwards rather than forwards as seen before. In the proof body we - encounter the @{command fix}-@{command assume}-@{command show} - outline of nested sub-proofs that is typical for Isar. The final - @{command show} is like @{command have} followed by an additional - refinement of the enclosing claim, using the rule derived from the - proof body. - - \medskip The next example involves @{term "\\"}, which can be - characterized as the set of all @{term "x"} such that @{prop "\A. x - \ A \ A \ \"}. The elimination rule for @{prop "x \ \\"} does - not mention @{text "\"} and @{text "\"} at all, but admits to obtain - directly a local @{term "A"} such that @{prop "x \ A"} and @{prop "A - \ \"} hold. This corresponds to the following Isar proof and - inference rule, respectively: -*} - -text_raw {*\medskip\begin{minipage}{0.6\textwidth}*} - -(*<*) -notepad -begin -(*>*) - assume "x \ \\" - then have C - proof - fix A - assume "x \ A" and "A \ \" - show C sorry %noproof - qed -(*<*) -end -(*>*) - -text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*} - -text {* - \infer{@{prop "C"}}{@{prop "x \ \\"} & \infer*{@{prop "C"}~}{@{text "[A][x \ A, A \ \]"}}} -*} - -text_raw {*\end{minipage}*} - -text {* - \medskip\noindent Although the Isar proof follows the natural - deduction rule closely, the text reads not as natural as - anticipated. There is a double occurrence of an arbitrary - conclusion @{prop "C"}, which represents the final result, but is - irrelevant for now. This issue arises for any elimination rule - involving local parameters. Isar provides the derived language - element @{command obtain}, which is able to perform the same - elimination proof more conveniently: -*} - -(*<*) -notepad -begin -(*>*) - assume "x \ \\" - then obtain A where "x \ A" and "A \ \" .. -(*<*) -end -(*>*) - -text {* - \noindent Here we avoid to mention the final conclusion @{prop "C"} - and return to plain forward reasoning. The rule involved in the - ``@{command ".."}'' proof is the same as before. -*} - - -section {* The Pure framework \label{sec:framework-pure} *} - -text {* - The Pure logic \cite{paulson-found,paulson700} is an intuitionistic - fragment of higher-order logic \cite{church40}. In type-theoretic - parlance, there are three levels of @{text "\"}-calculus with - corresponding arrows @{text "\"}/@{text "\"}/@{text "\"}: - - \medskip - \begin{tabular}{ll} - @{text "\ \ \"} & syntactic function space (terms depending on terms) \\ - @{text "\x. B(x)"} & universal quantification (proofs depending on terms) \\ - @{text "A \ B"} & implication (proofs depending on proofs) \\ - \end{tabular} - \medskip - - \noindent Here only the types of syntactic terms, and the - propositions of proof terms have been shown. The @{text - "\"}-structure of proofs can be recorded as an optional feature of - the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but - the formal system can never depend on them due to \emph{proof - irrelevance}. - - On top of this most primitive layer of proofs, Pure implements a - generic calculus for nested natural deduction rules, similar to - \cite{Schroeder-Heister:1984}. Here object-logic inferences are - internalized as formulae over @{text "\"} and @{text "\"}. - Combining such rule statements may involve higher-order unification - \cite{paulson-natural}. -*} - - -subsection {* Primitive inferences *} - -text {* - Term syntax provides explicit notation for abstraction @{text "\x :: - \. b(x)"} and application @{text "b a"}, while types are usually - implicit thanks to type-inference; terms of type @{text "prop"} are - called propositions. Logical statements are composed via @{text "\x - :: \. B(x)"} and @{text "A \ B"}. Primitive reasoning operates on - judgments of the form @{text "\ \ \"}, with standard introduction - and elimination rules for @{text "\"} and @{text "\"} that refer to - fixed parameters @{text "x\<^isub>1, \, x\<^isub>m"} and hypotheses - @{text "A\<^isub>1, \, A\<^isub>n"} from the context @{text "\"}; - the corresponding proof terms are left implicit. The subsequent - inference rules define @{text "\ \ \"} inductively, relative to a - collection of axioms: - - \[ - \infer{@{text "\ A"}}{(@{text "A"} \text{~axiom})} - \qquad - \infer{@{text "A \ A"}}{} - \] - - \[ - \infer{@{text "\ \ \x. B(x)"}}{@{text "\ \ B(x)"} & @{text "x \ \"}} - \qquad - \infer{@{text "\ \ B(a)"}}{@{text "\ \ \x. B(x)"}} - \] - - \[ - \infer{@{text "\ - A \ A \ B"}}{@{text "\ \ B"}} - \qquad - \infer{@{text "\\<^sub>1 \ \\<^sub>2 \ B"}}{@{text "\\<^sub>1 \ A \ B"} & @{text "\\<^sub>2 \ A"}} - \] - - Furthermore, Pure provides a built-in equality @{text "\ :: \ \ \ \ - prop"} with axioms for reflexivity, substitution, extensionality, - and @{text "\\\"}-conversion on @{text "\"}-terms. - - \medskip An object-logic introduces another layer on top of Pure, - e.g.\ with types @{text "i"} for individuals and @{text "o"} for - propositions, term constants @{text "Trueprop :: o \ prop"} as - (implicit) derivability judgment and connectives like @{text "\ :: o - \ o \ o"} or @{text "\ :: (i \ o) \ o"}, and axioms for object-level - rules such as @{text "conjI: A \ B \ A \ B"} or @{text "allI: (\x. B - x) \ \x. B x"}. Derived object rules are represented as theorems of - Pure. After the initial object-logic setup, further axiomatizations - are usually avoided; plain definitions and derived principles are - used exclusively. -*} - - -subsection {* Reasoning with rules \label{sec:framework-resolution} *} - -text {* - Primitive inferences mostly serve foundational purposes. The main - reasoning mechanisms of Pure operate on nested natural deduction - rules expressed as formulae, using @{text "\"} to bind local - parameters and @{text "\"} to express entailment. Multiple - parameters and premises are represented by repeating these - connectives in a right-associative manner. - - Since @{text "\"} and @{text "\"} commute thanks to the theorem - @{prop "(A \ (\x. B x)) \ (\x. A \ B x)"}, we may assume w.l.o.g.\ - that rule statements always observe the normal form where - quantifiers are pulled in front of implications at each level of - nesting. This means that any Pure proposition may be presented as a - \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the - form @{text "\x\<^isub>1 \ x\<^isub>m. H\<^isub>1 \ \ H\<^isub>n \ - A"} for @{text "m, n \ 0"}, and @{text "A"} atomic, and @{text - "H\<^isub>1, \, H\<^isub>n"} being recursively of the same format. - Following the convention that outermost quantifiers are implicit, - Horn clauses @{text "A\<^isub>1 \ \ A\<^isub>n \ A"} are a special - case of this. - - For example, @{text "\"}-introduction rule encountered before is - represented as a Pure theorem as follows: - \[ - @{text "IntI:"}~@{prop "x \ A \ x \ B \ x \ A \ B"} - \] - - \noindent This is a plain Horn clause, since no further nesting on - the left is involved. The general @{text "\"}-introduction - corresponds to a Hereditary Harrop Formula with one additional level - of nesting: - \[ - @{text "InterI:"}~@{prop "(\A. A \ \ \ x \ A) \ x \ \\"} - \] - - \medskip Goals are also represented as rules: @{text "A\<^isub>1 \ - \ A\<^isub>n \ C"} states that the sub-goals @{text "A\<^isub>1, \, - A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the - goal is finished. To allow @{text "C"} being a rule statement - itself, we introduce the protective marker @{text "# :: prop \ - prop"}, which is defined as identity and hidden from the user. We - initialize and finish goal states as follows: - - \[ - \begin{array}{c@ {\qquad}c} - \infer[(@{inference_def init})]{@{text "C \ #C"}}{} & - \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}} - \end{array} - \] - - \noindent Goal states are refined in intermediate proof steps until - a finished form is achieved. Here the two main reasoning principles - are @{inference resolution}, for back-chaining a rule against a - sub-goal (replacing it by zero or more sub-goals), and @{inference - assumption}, for solving a sub-goal (finding a short-circuit with - local assumptions). Below @{text "\<^vec>x"} stands for @{text - "x\<^isub>1, \, x\<^isub>n"} (@{text "n \ 0"}). - - \[ - \infer[(@{inference_def resolution})] - {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (\<^vec>a \<^vec>x))\ \ C\"}} - {\begin{tabular}{rl} - @{text "rule:"} & - @{text "\<^vec>A \<^vec>a \ B \<^vec>a"} \\ - @{text "goal:"} & - @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ - @{text "goal unifier:"} & - @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ - \end{tabular}} - \] - - \medskip - - \[ - \infer[(@{inference_def assumption})]{@{text "C\"}} - {\begin{tabular}{rl} - @{text "goal:"} & - @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C"} \\ - @{text "assm unifier:"} & @{text "A\ = H\<^sub>i\"}~~\text{(for some~@{text "H\<^sub>i"})} \\ - \end{tabular}} - \] - - The following trace illustrates goal-oriented reasoning in - Isabelle/Pure: - - {\footnotesize - \medskip - \begin{tabular}{r@ {\quad}l} - @{text "(A \ B \ B \ A) \ #(A \ B \ B \ A)"} & @{text "(init)"} \\ - @{text "(A \ B \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution B \ A \ B \ A)"} \\ - @{text "(A \ B \ A \ B) \ (A \ B \ A) \ #\"} & @{text "(resolution A \ B \ B)"} \\ - @{text "(A \ B \ A) \ #\"} & @{text "(assumption)"} \\ - @{text "(A \ B \ A \ B) \ #\"} & @{text "(resolution A \ B \ A)"} \\ - @{text "#\"} & @{text "(assumption)"} \\ - @{text "A \ B \ B \ A"} & @{text "(finish)"} \\ - \end{tabular} - \medskip - } - - Compositions of @{inference assumption} after @{inference - resolution} occurs quite often, typically in elimination steps. - Traditional Isabelle tactics accommodate this by a combined - @{inference_def elim_resolution} principle. In contrast, Isar uses - a slightly more refined combination, where the assumptions to be - closed are marked explicitly, using again the protective marker - @{text "#"}: - - \[ - \infer[(@{inference refinement})] - {@{text "(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>G' (\<^vec>a \<^vec>x))\ \ C\"}} - {\begin{tabular}{rl} - @{text "sub\proof:"} & - @{text "\<^vec>G \<^vec>a \ B \<^vec>a"} \\ - @{text "goal:"} & - @{text "(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C"} \\ - @{text "goal unifier:"} & - @{text "(\\<^vec>x. B (\<^vec>a \<^vec>x))\ = B'\"} \\ - @{text "assm unifiers:"} & - @{text "(\\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\ = #H\<^sub>i\"} \\ - & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\ - \end{tabular}} - \] - - \noindent Here the @{text "sub\proof"} rule stems from the - main @{command fix}-@{command assume}-@{command show} outline of - Isar (cf.\ \secref{sec:framework-subproof}): each assumption - indicated in the text results in a marked premise @{text "G"} above. - The marking enforces resolution against one of the sub-goal's - premises. Consequently, @{command fix}-@{command assume}-@{command - show} enables to fit the result of a sub-proof quite robustly into a - pending sub-goal, while maintaining a good measure of flexibility. -*} - - -section {* The Isar proof language \label{sec:framework-isar} *} - -text {* - Structured proofs are presented as high-level expressions for - composing entities of Pure (propositions, facts, and goals). The - Isar proof language allows to organize reasoning within the - underlying rule calculus of Pure, but Isar is not another logical - calculus! - - Isar is an exercise in sound minimalism. Approximately half of the - language is introduced as primitive, the rest defined as derived - concepts. The following grammar describes the core language - (category @{text "proof"}), which is embedded into theory - specification elements such as @{command theorem}; see also - \secref{sec:framework-stmt} for the separate category @{text - "statement"}. - - \medskip - \begin{tabular}{rcl} - @{text "theory\stmt"} & = & @{command "theorem"}~@{text "statement proof |"}~~@{command "definition"}~@{text "\ | \"} \\[1ex] - - @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex] - - @{text prfx} & = & @{command "using"}~@{text "facts"} \\ - & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ - - @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ - & @{text "|"} & @{command "next"} \\ - & @{text "|"} & @{command "note"}~@{text "name = facts"} \\ - & @{text "|"} & @{command "let"}~@{text "term = term"} \\ - & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ - & @{text "|"} & @{command assume}~@{text "\inference\ name: props"} \\ - & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ - @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ - & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ - \end{tabular} - - \medskip Simultaneous propositions or facts may be separated by the - @{keyword "and"} keyword. - - \medskip The syntax for terms and propositions is inherited from - Pure (and the object-logic). A @{text "pattern"} is a @{text - "term"} with schematic variables, to be bound by higher-order - matching. - - \medskip Facts may be referenced by name or proposition. For - example, the result of ``@{command have}~@{text "a: A \proof\"}'' - becomes available both as @{text "a"} and - \isacharbackquoteopen@{text "A"}\isacharbackquoteclose. Moreover, - fact expressions may involve attributes that modify either the - theorem or the background context. For example, the expression - ``@{text "a [OF b]"}'' refers to the composition of two facts - according to the @{inference resolution} inference of - \secref{sec:framework-resolution}, while ``@{text "a [intro]"}'' - declares a fact as introduction rule in the context. - - The special fact called ``@{fact this}'' always refers to the last - result, as produced by @{command note}, @{command assume}, @{command - have}, or @{command show}. Since @{command note} occurs - frequently together with @{command then} we provide some - abbreviations: - - \medskip - \begin{tabular}{rcl} - @{command from}~@{text a} & @{text "\"} & @{command note}~@{text a}~@{command then} \\ - @{command with}~@{text a} & @{text "\"} & @{command from}~@{text "a \ this"} \\ - \end{tabular} - \medskip - - The @{text "method"} category is essentially a parameter and may be - populated later. Methods use the facts indicated by @{command - "then"} or @{command using}, and then operate on the goal state. - Some basic methods are predefined: ``@{method "-"}'' leaves the goal - unchanged, ``@{method this}'' applies the facts as rules to the - goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the - result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}'' - refer to @{inference resolution} of - \secref{sec:framework-resolution}). The secondary arguments to - ``@{method (Pure) rule}'' may be specified explicitly as in ``@{text "(rule - a)"}'', or picked from the context. In the latter case, the system - first tries rules declared as @{attribute (Pure) elim} or - @{attribute (Pure) dest}, followed by those declared as @{attribute - (Pure) intro}. - - The default method for @{command proof} is ``@{method (Pure) rule}'' - (arguments picked from the context), for @{command qed} it is - ``@{method "-"}''. Further abbreviations for terminal proof steps - are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for - ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text - "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command - "by"}~@{method (Pure) rule}, and ``@{command "."}'' for ``@{command - "by"}~@{method this}''. The @{command unfolding} element operates - directly on the current facts and goal by applying equalities. - - \medskip Block structure can be indicated explicitly by ``@{command - "{"}~@{text "\"}~@{command "}"}'', although the body of a sub-proof - already involves implicit nesting. In any case, @{command next} - jumps into the next section of a block, i.e.\ it acts like closing - an implicit block scope and opening another one; there is no direct - correspondence to subgoals here. - - The remaining elements @{command fix} and @{command assume} build up - a local context (see \secref{sec:framework-context}), while - @{command show} refines a pending sub-goal by the rule resulting - from a nested sub-proof (see \secref{sec:framework-subproof}). - Further derived concepts will support calculational reasoning (see - \secref{sec:framework-calc}). -*} - - -subsection {* Context elements \label{sec:framework-context} *} - -text {* - In judgments @{text "\ \ \"} of the primitive framework, @{text "\"} - essentially acts like a proof context. Isar elaborates this idea - towards a higher-level notion, with additional information for - type-inference, term abbreviations, local facts, hypotheses etc. - - The element @{command fix}~@{text "x :: \"} declares a local - parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in - results exported from the context, @{text "x"} may become anything. - The @{command assume}~@{text "\inference\"} element provides a - general interface to hypotheses: ``@{command assume}~@{text - "\inference\ A"}'' produces @{text "A \ A"} locally, while the - included inference tells how to discharge @{text A} from results - @{text "A \ B"} later on. There is no user-syntax for @{text - "\inference\"}, i.e.\ it may only occur internally when derived - commands are defined in ML. - - At the user-level, the default inference for @{command assume} is - @{inference discharge} as given below. The additional variants - @{command presume} and @{command def} are defined as follows: - - \medskip - \begin{tabular}{rcl} - @{command presume}~@{text A} & @{text "\"} & @{command assume}~@{text "\weak\discharge\ A"} \\ - @{command def}~@{text "x \ a"} & @{text "\"} & @{command fix}~@{text x}~@{command assume}~@{text "\expansion\ x \ a"} \\ - \end{tabular} - \medskip - - \[ - \infer[(@{inference_def discharge})]{@{text "\\ - A \ #A \ B"}}{@{text "\\ \ B"}} - \] - \[ - \infer[(@{inference_def "weak\discharge"})]{@{text "\\ - A \ A \ B"}}{@{text "\\ \ B"}} - \] - \[ - \infer[(@{inference_def expansion})]{@{text "\\ - (x \ a) \ B a"}}{@{text "\\ \ B x"}} - \] - - \medskip Note that @{inference discharge} and @{inference - "weak\discharge"} differ in the marker for @{prop A}, which is - relevant when the result of a @{command fix}-@{command - assume}-@{command show} outline is composed with a pending goal, - cf.\ \secref{sec:framework-subproof}. - - The most interesting derived context element in Isar is @{command - obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized - elimination steps in a purely forward manner. The @{command obtain} - command takes a specification of parameters @{text "\<^vec>x"} and - assumptions @{text "\<^vec>A"} to be added to the context, together - with a proof of a case rule stating that this extension is - conservative (i.e.\ may be removed from closed results later on): - - \medskip - \begin{tabular}{l} - @{text "\facts\"}~~@{command obtain}~@{text "\<^vec>x \ \<^vec>A \<^vec>x \proof\ \"} \\[0.5ex] - \quad @{command have}~@{text "case: \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis\"} \\ - \quad @{command proof}~@{method "-"} \\ - \qquad @{command fix}~@{text thesis} \\ - \qquad @{command assume}~@{text "[intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis"} \\ - \qquad @{command show}~@{text thesis}~@{command using}~@{text "\facts\ \proof\"} \\ - \quad @{command qed} \\ - \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\elimination case\ \<^vec>A \<^vec>x"} \\ - \end{tabular} - \medskip - - \[ - \infer[(@{inference elimination})]{@{text "\ \ B"}}{ - \begin{tabular}{rl} - @{text "case:"} & - @{text "\ \ \thesis. (\\<^vec>x. \<^vec>A \<^vec>x \ thesis) \ thesis"} \\[0.2ex] - @{text "result:"} & - @{text "\ \ \<^vec>A \<^vec>y \ B"} \\[0.2ex] - \end{tabular}} - \] - - \noindent Here the name ``@{text thesis}'' is a specific convention - for an arbitrary-but-fixed proposition; in the primitive natural - deduction rules shown before we have occasionally used @{text C}. - The whole statement of ``@{command obtain}~@{text x}~@{keyword - "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"} - may be assumed for some arbitrary-but-fixed @{text "x"}. Also note - that ``@{command obtain}~@{text "A \ B"}'' without parameters - is similar to ``@{command have}~@{text "A \ B"}'', but the - latter involves multiple sub-goals. - - \medskip The subsequent Isar proof texts explain all context - elements introduced above using the formal proof language itself. - After finishing a local proof within a block, we indicate the - exported result via @{command note}. -*} - -(*<*) -theorem True -proof -(*>*) - txt_raw {* \begin{minipage}[t]{0.45\textwidth} *} - { - fix x - have "B x" sorry %noproof - } - note `\x. B x` - txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) - { - assume A - have B sorry %noproof - } - note `A \ B` - txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) - { - def x \ a - have "B x" sorry %noproof - } - note `B a` - txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*) - { - obtain x where "A x" sorry %noproof - have B sorry %noproof - } - note `B` - txt_raw {* \end{minipage} *} -(*<*) -qed -(*>*) - -text {* - \bigskip\noindent This illustrates the meaning of Isar context - elements without goals getting in between. -*} - -subsection {* Structured statements \label{sec:framework-stmt} *} - -text {* - The category @{text "statement"} of top-level theorem specifications - is defined as follows: - - \medskip - \begin{tabular}{rcl} - @{text "statement"} & @{text "\"} & @{text "name: props \ \"} \\ - & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex] - - @{text "context"} & @{text "\"} & @{text "\ vars \ \"} \\ - & @{text "|"} & @{text "\ name: props \ \"} \\ - - @{text "conclusion"} & @{text "\"} & @{text "\ name: props \ \"} \\ - & @{text "|"} & @{text "\ vars \ \ \ name: props \ \"} \\ - & & \quad @{text "\ \"} \\ - \end{tabular} - - \medskip\noindent A simple @{text "statement"} consists of named - propositions. The full form admits local context elements followed - by the actual conclusions, such as ``@{keyword "fixes"}~@{text - x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B - x"}''. The final result emerges as a Pure rule after discharging - the context: @{prop "\x. A x \ B x"}. - - The @{keyword "obtains"} variant is another abbreviation defined - below; unlike @{command obtain} (cf.\ - \secref{sec:framework-context}) there may be several ``cases'' - separated by ``@{text "\"}'', each consisting of several - parameters (@{text "vars"}) and several premises (@{text "props"}). - This specifies multi-branch elimination rules. - - \medskip - \begin{tabular}{l} - @{text "\ \<^vec>x \ \<^vec>A \<^vec>x \ \ \"} \\[0.5ex] - \quad @{text "\ thesis"} \\ - \quad @{text "\ [intro]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis \ \"} \\ - \quad @{text "\ thesis"} \\ - \end{tabular} - \medskip - - Presenting structured statements in such an ``open'' format usually - simplifies the subsequent proof, because the outer structure of the - problem is already laid out directly. E.g.\ consider the following - canonical patterns for @{text "\"} and @{text "\"}, - respectively: -*} - -text_raw {*\begin{minipage}{0.5\textwidth}*} - -theorem - fixes x and y - assumes "A x" and "B y" - shows "C x y" -proof - - from `A x` and `B y` - show "C x y" sorry %noproof -qed - -text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} - -theorem - obtains x and y - where "A x" and "B y" -proof - - have "A a" and "B b" sorry %noproof - then show thesis .. -qed - -text_raw {*\end{minipage}*} - -text {* - \medskip\noindent Here local facts \isacharbackquoteopen@{text "A - x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B - y"}\isacharbackquoteclose\ are referenced immediately; there is no - need to decompose the logical rule structure again. In the second - proof the final ``@{command then}~@{command show}~@{text - thesis}~@{command ".."}'' involves the local rule case @{text "\x - y. A x \ B y \ thesis"} for the particular instance of terms @{text - "a"} and @{text "b"} produced in the body. -*} - - -subsection {* Structured proof refinement \label{sec:framework-subproof} *} - -text {* - By breaking up the grammar for the Isar proof language, we may - understand a proof text as a linear sequence of individual proof - commands. These are interpreted as transitions of the Isar virtual - machine (Isar/VM), which operates on a block-structured - configuration in single steps. This allows users to write proof - texts in an incremental manner, and inspect intermediate - configurations for debugging. - - The basic idea is analogous to evaluating algebraic expressions on a - stack machine: @{text "(a + b) \ c"} then corresponds to a sequence - of single transitions for each symbol @{text "(, a, +, b, ), \, c"}. - In Isar the algebraic values are facts or goals, and the operations - are inferences. - - \medskip The Isar/VM state maintains a stack of nodes, each node - contains the local proof context, the linguistic mode, and a pending - goal (optional). The mode determines the type of transition that - may be performed next, it essentially alternates between forward and - backward reasoning, with an intermediate stage for chained facts - (see \figref{fig:isar-vm}). - - \begin{figure}[htb] - \begin{center} - \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm} - \end{center} - \caption{Isar/VM modes}\label{fig:isar-vm} - \end{figure} - - For example, in @{text "state"} mode Isar acts like a mathematical - scratch-pad, accepting declarations like @{command fix}, @{command - assume}, and claims like @{command have}, @{command show}. A goal - statement changes the mode to @{text "prove"}, which means that we - may now refine the problem via @{command unfolding} or @{command - proof}. Then we are again in @{text "state"} mode of a proof body, - which may issue @{command show} statements to solve pending - sub-goals. A concluding @{command qed} will return to the original - @{text "state"} mode one level upwards. The subsequent Isar/VM - trace indicates block structure, linguistic mode, goal state, and - inferences: -*} - -text_raw {* \begingroup\footnotesize *} -(*<*)notepad begin -(*>*) - txt_raw {* \begin{minipage}[t]{0.18\textwidth} *} - have "A \ B" - proof - assume A - show B - sorry %noproof - qed - txt_raw {* \end{minipage}\quad -\begin{minipage}[t]{0.06\textwidth} -@{text "begin"} \\ -\\ -\\ -@{text "begin"} \\ -@{text "end"} \\ -@{text "end"} \\ -\end{minipage} -\begin{minipage}[t]{0.08\textwidth} -@{text "prove"} \\ -@{text "state"} \\ -@{text "state"} \\ -@{text "prove"} \\ -@{text "state"} \\ -@{text "state"} \\ -\end{minipage}\begin{minipage}[t]{0.35\textwidth} -@{text "(A \ B) \ #(A \ B)"} \\ -@{text "(A \ B) \ #(A \ B)"} \\ -\\ -\\ -@{text "#(A \ B)"} \\ -@{text "A \ B"} \\ -\end{minipage}\begin{minipage}[t]{0.4\textwidth} -@{text "(init)"} \\ -@{text "(resolution impI)"} \\ -\\ -\\ -@{text "(refinement #A \ B)"} \\ -@{text "(finish)"} \\ -\end{minipage} *} -(*<*) -end -(*>*) -text_raw {* \endgroup *} - -text {* - \noindent Here the @{inference refinement} inference from - \secref{sec:framework-resolution} mediates composition of Isar - sub-proofs nicely. Observe that this principle incorporates some - degree of freedom in proof composition. In particular, the proof - body allows parameters and assumptions to be re-ordered, or commuted - according to Hereditary Harrop Form. Moreover, context elements - that are not used in a sub-proof may be omitted altogether. For - example: -*} - -text_raw {*\begin{minipage}{0.5\textwidth}*} - -(*<*) -notepad -begin -(*>*) - have "\x y. A x \ B y \ C x y" - proof - - fix x and y - assume "A x" and "B y" - show "C x y" sorry %noproof - qed - -txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} - -(*<*) -next -(*>*) - have "\x y. A x \ B y \ C x y" - proof - - fix x assume "A x" - fix y assume "B y" - show "C x y" sorry %noproof - qed - -txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*} - -(*<*) -next -(*>*) - have "\x y. A x \ B y \ C x y" - proof - - fix y assume "B y" - fix x assume "A x" - show "C x y" sorry - qed - -txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*} -(*<*) -next -(*>*) - have "\x y. A x \ B y \ C x y" - proof - - fix y assume "B y" - fix x - show "C x y" sorry - qed -(*<*) -end -(*>*) - -text_raw {*\end{minipage}*} - -text {* - \medskip\noindent Such ``peephole optimizations'' of Isar texts are - practically important to improve readability, by rearranging - contexts elements according to the natural flow of reasoning in the - body, while still observing the overall scoping rules. - - \medskip This illustrates the basic idea of structured proof - processing in Isar. The main mechanisms are based on natural - deduction rule composition within the Pure framework. In - particular, there are no direct operations on goal states within the - proof body. Moreover, there is no hidden automated reasoning - involved, just plain unification. -*} - - -subsection {* Calculational reasoning \label{sec:framework-calc} *} - -text {* - The existing Isar infrastructure is sufficiently flexible to support - calculational reasoning (chains of transitivity steps) as derived - concept. The generic proof elements introduced below depend on - rules declared as @{attribute trans} in the context. It is left to - the object-logic to provide a suitable rule collection for mixed - relations of @{text "="}, @{text "<"}, @{text "\"}, @{text "\"}, - @{text "\"} etc. Due to the flexibility of rule composition - (\secref{sec:framework-resolution}), substitution of equals by - equals is covered as well, even substitution of inequalities - involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD} - and \cite{Bauer-Wenzel:2001}. - - The generic calculational mechanism is based on the observation that - rules such as @{text "trans:"}~@{prop "x = y \ y = z \ x = z"} - proceed from the premises towards the conclusion in a deterministic - fashion. Thus we may reason in forward mode, feeding intermediate - results into rules selected from the context. The course of - reasoning is organized by maintaining a secondary fact called - ``@{fact calculation}'', apart from the primary ``@{fact this}'' - already provided by the Isar primitives. In the definitions below, - @{attribute OF} refers to @{inference resolution} - (\secref{sec:framework-resolution}) with multiple rule arguments, - and @{text "trans"} represents to a suitable rule from the context: - - \begin{matharray}{rcl} - @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\ - @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex] - @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\ - \end{matharray} - - \noindent The start of a calculation is determined implicitly in the - text: here @{command also} sets @{fact calculation} to the current - result; any subsequent occurrence will update @{fact calculation} by - combination with the next result and a transitivity rule. The - calculational sequence is concluded via @{command finally}, where - the final result is exposed for use in a concluding claim. - - Here is a canonical proof pattern, using @{command have} to - establish the intermediate results: -*} - -(*<*) -notepad -begin -(*>*) - have "a = b" sorry - also have "\ = c" sorry - also have "\ = d" sorry - finally have "a = d" . -(*<*) -end -(*>*) - -text {* - \noindent The term ``@{text "\"}'' above is a special abbreviation - provided by the Isabelle/Isar syntax layer: it statically refers to - the right-hand side argument of the previous statement given in the - text. Thus it happens to coincide with relevant sub-expressions in - the calculational chain, but the exact correspondence is dependent - on the transitivity rules being involved. - - \medskip Symmetry rules such as @{prop "x = y \ y = x"} are like - transitivities with only one premise. Isar maintains a separate - rule collection declared via the @{attribute sym} attribute, to be - used in fact expressions ``@{text "a [symmetric]"}'', or single-step - proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command - have}~@{text "y = x"}~@{command ".."}''. -*} - -end \ No newline at end of file diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Generic.thy --- a/doc-src/IsarRef/Thy/Generic.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1279 +0,0 @@ -theory Generic -imports Base Main -begin - -chapter {* Generic tools and packages \label{ch:gen-tools} *} - -section {* Configuration options \label{sec:config} *} - -text {* Isabelle/Pure maintains a record of named configuration - options within the theory or proof context, with values of type - @{ML_type bool}, @{ML_type int}, @{ML_type real}, or @{ML_type - string}. Tools may declare options in ML, and then refer to these - values (relative to the context). Thus global reference variables - are easily avoided. The user may change the value of a - configuration option by means of an associated attribute of the same - name. This form of context declaration works particularly well with - commands such as @{command "declare"} or @{command "using"} like - this: -*} - -declare [[show_main_goal = false]] - -notepad -begin - note [[show_main_goal = true]] -end - -text {* For historical reasons, some tools cannot take the full proof - context into account and merely refer to the background theory. - This is accommodated by configuration options being declared as - ``global'', which may not be changed within a local context. - - \begin{matharray}{rcll} - @{command_def "print_configs"} & : & @{text "context \"} \\ - \end{matharray} - - @{rail " - @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? - "} - - \begin{description} - - \item @{command "print_configs"} prints the available configuration - options, with names, types, and current values. - - \item @{text "name = value"} as an attribute expression modifies the - named option, with the syntax of the value depending on the option's - type. For @{ML_type bool} the default value is @{text true}. Any - attempt to change a global option in a local context is ignored. - - \end{description} -*} - - -section {* Basic proof tools *} - -subsection {* Miscellaneous methods and attributes \label{sec:misc-meth-att} *} - -text {* - \begin{matharray}{rcl} - @{method_def unfold} & : & @{text method} \\ - @{method_def fold} & : & @{text method} \\ - @{method_def insert} & : & @{text method} \\[0.5ex] - @{method_def erule}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def drule}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def frule}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def intro} & : & @{text method} \\ - @{method_def elim} & : & @{text method} \\ - @{method_def succeed} & : & @{text method} \\ - @{method_def fail} & : & @{text method} \\ - \end{matharray} - - @{rail " - (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thmrefs} - ; - (@@{method erule} | @@{method drule} | @@{method frule}) - ('(' @{syntax nat} ')')? @{syntax thmrefs} - ; - (@@{method intro} | @@{method elim}) @{syntax thmrefs}? - "} - - \begin{description} - - \item @{method unfold}~@{text "a\<^sub>1 \ a\<^sub>n"} and @{method fold}~@{text - "a\<^sub>1 \ a\<^sub>n"} expand (or fold back) the given definitions throughout - all goals; any chained facts provided are inserted into the goal and - subject to rewriting as well. - - \item @{method insert}~@{text "a\<^sub>1 \ a\<^sub>n"} inserts theorems as facts - into all goals of the proof state. Note that current facts - indicated for forward chaining are ignored. - - \item @{method erule}~@{text "a\<^sub>1 \ a\<^sub>n"}, @{method - drule}~@{text "a\<^sub>1 \ a\<^sub>n"}, and @{method frule}~@{text - "a\<^sub>1 \ a\<^sub>n"} are similar to the basic @{method rule} - method (see \secref{sec:pure-meth-att}), but apply rules by - elim-resolution, destruct-resolution, and forward-resolution, - respectively \cite{isabelle-implementation}. The optional natural - number argument (default 0) specifies additional assumption steps to - be performed here. - - Note that these methods are improper ones, mainly serving for - experimentation and tactic script emulation. Different modes of - basic rule application are usually expressed in Isar at the proof - language level, rather than via implicit proof state manipulations. - For example, a proper single-step elimination would be done using - the plain @{method rule} method, with forward chaining of current - facts. - - \item @{method intro} and @{method elim} repeatedly refine some goal - by intro- or elim-resolution, after having inserted any chained - facts. Exactly the rules given as arguments are taken into account; - this allows fine-tuned decomposition of a proof problem, in contrast - to common automated tools. - - \item @{method succeed} yields a single (unchanged) result; it is - the identity of the ``@{text ","}'' method combinator (cf.\ - \secref{sec:proof-meth}). - - \item @{method fail} yields an empty result sequence; it is the - identity of the ``@{text "|"}'' method combinator (cf.\ - \secref{sec:proof-meth}). - - \end{description} - - \begin{matharray}{rcl} - @{attribute_def tagged} & : & @{text attribute} \\ - @{attribute_def untagged} & : & @{text attribute} \\[0.5ex] - @{attribute_def THEN} & : & @{text attribute} \\ - @{attribute_def unfolded} & : & @{text attribute} \\ - @{attribute_def folded} & : & @{text attribute} \\ - @{attribute_def abs_def} & : & @{text attribute} \\[0.5ex] - @{attribute_def rotated} & : & @{text attribute} \\ - @{attribute_def (Pure) elim_format} & : & @{text attribute} \\ - @{attribute_def standard}@{text "\<^sup>*"} & : & @{text attribute} \\ - @{attribute_def no_vars}@{text "\<^sup>*"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{attribute tagged} @{syntax name} @{syntax name} - ; - @@{attribute untagged} @{syntax name} - ; - @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thmref} - ; - (@@{attribute unfolded} | @@{attribute folded}) @{syntax thmrefs} - ; - @@{attribute rotated} @{syntax int}? - "} - - \begin{description} - - \item @{attribute tagged}~@{text "name value"} and @{attribute - untagged}~@{text name} add and remove \emph{tags} of some theorem. - Tags may be any list of string pairs that serve as formal comment. - The first string is considered the tag name, the second its value. - Note that @{attribute untagged} removes any tags of the same name. - - \item @{attribute THEN}~@{text a} composes rules by resolution; it - resolves with the first premise of @{text a} (an alternative - position may be also specified). See also @{ML_op "RS"} in - \cite{isabelle-implementation}. - - \item @{attribute unfolded}~@{text "a\<^sub>1 \ a\<^sub>n"} and @{attribute - folded}~@{text "a\<^sub>1 \ a\<^sub>n"} expand and fold back again the given - definitions throughout a rule. - - \item @{attribute abs_def} turns an equation of the form @{prop "f x - y \ t"} into @{prop "f \ \x y. t"}, which ensures that @{method - simp} or @{method unfold} steps always expand it. This also works - for object-logic equality. - - \item @{attribute rotated}~@{text n} rotate the premises of a - theorem by @{text n} (default 1). - - \item @{attribute (Pure) elim_format} turns a destruction rule into - elimination rule format, by resolving with the rule @{prop "PROP A \ - (PROP A \ PROP B) \ PROP B"}. - - Note that the Classical Reasoner (\secref{sec:classical}) provides - its own version of this operation. - - \item @{attribute standard} puts a theorem into the standard form of - object-rules at the outermost theory level. Note that this - operation violates the local proof context (including active - locales). - - \item @{attribute no_vars} replaces schematic variables by free - ones; this is mainly for tuning output of pretty printed theorems. - - \end{description} -*} - - -subsection {* Low-level equational reasoning *} - -text {* - \begin{matharray}{rcl} - @{method_def subst} & : & @{text method} \\ - @{method_def hypsubst} & : & @{text method} \\ - @{method_def split} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method subst} ('(' 'asm' ')')? \\ ('(' (@{syntax nat}+) ')')? @{syntax thmref} - ; - @@{method split} @{syntax thmrefs} - "} - - These methods provide low-level facilities for equational reasoning - that are intended for specialized applications only. Normally, - single step calculations would be performed in a structured text - (see also \secref{sec:calculation}), while the Simplifier methods - provide the canonical way for automated normalization (see - \secref{sec:simplifier}). - - \begin{description} - - \item @{method subst}~@{text eq} performs a single substitution step - using rule @{text eq}, which may be either a meta or object - equality. - - \item @{method subst}~@{text "(asm) eq"} substitutes in an - assumption. - - \item @{method subst}~@{text "(i \ j) eq"} performs several - substitutions in the conclusion. The numbers @{text i} to @{text j} - indicate the positions to substitute at. Positions are ordered from - the top of the term tree moving down from left to right. For - example, in @{text "(a + b) + (c + d)"} there are three positions - where commutativity of @{text "+"} is applicable: 1 refers to @{text - "a + b"}, 2 to the whole term, and 3 to @{text "c + d"}. - - If the positions in the list @{text "(i \ j)"} are non-overlapping - (e.g.\ @{text "(2 3)"} in @{text "(a + b) + (c + d)"}) you may - assume all substitutions are performed simultaneously. Otherwise - the behaviour of @{text subst} is not specified. - - \item @{method subst}~@{text "(asm) (i \ j) eq"} performs the - substitutions in the assumptions. The positions refer to the - assumptions in order from left to right. For example, given in a - goal of the form @{text "P (a + b) \ P (c + d) \ \"}, position 1 of - commutativity of @{text "+"} is the subterm @{text "a + b"} and - position 2 is the subterm @{text "c + d"}. - - \item @{method hypsubst} performs substitution using some - assumption; this only works for equations of the form @{text "x = - t"} where @{text x} is a free or bound variable. - - \item @{method split}~@{text "a\<^sub>1 \ a\<^sub>n"} performs single-step case - splitting using the given rules. Splitting is performed in the - conclusion or some assumption of the subgoal, depending of the - structure of the rule. - - Note that the @{method simp} method already involves repeated - application of split rules as declared in the current context, using - @{attribute split}, for example. - - \end{description} -*} - - -subsection {* Further tactic emulations \label{sec:tactics} *} - -text {* - The following improper proof methods emulate traditional tactics. - These admit direct access to the goal state, which is normally - considered harmful! In particular, this may involve both numbered - goal addressing (default 1), and dynamic instantiation within the - scope of some subgoal. - - \begin{warn} - Dynamic instantiations refer to universally quantified parameters - of a subgoal (the dynamic context) rather than fixed variables and - term abbreviations of a (static) Isar context. - \end{warn} - - Tactic emulation methods, unlike their ML counterparts, admit - simultaneous instantiation from both dynamic and static contexts. - If names occur in both contexts goal parameters hide locally fixed - variables. Likewise, schematic variables refer to term - abbreviations, if present in the static context. Otherwise the - schematic variable is interpreted as a schematic variable and left - to be solved by unification with certain parts of the subgoal. - - Note that the tactic emulation proof methods in Isabelle/Isar are - consistently named @{text foo_tac}. Note also that variable names - occurring on left hand sides of instantiations must be preceded by a - question mark if they coincide with a keyword or contain dots. This - is consistent with the attribute @{attribute "where"} (see - \secref{sec:pure-meth-att}). - - \begin{matharray}{rcl} - @{method_def rule_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def erule_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def drule_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def frule_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def cut_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def thin_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def subgoal_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def rename_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def rotate_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def tactic}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def raw_tactic}@{text "\<^sup>*"} & : & @{text method} \\ - \end{matharray} - - @{rail " - (@@{method rule_tac} | @@{method erule_tac} | @@{method drule_tac} | - @@{method frule_tac} | @@{method cut_tac} | @@{method thin_tac}) @{syntax goal_spec}? \\ - ( dynamic_insts @'in' @{syntax thmref} | @{syntax thmrefs} ) - ; - @@{method subgoal_tac} @{syntax goal_spec}? (@{syntax prop} +) - ; - @@{method rename_tac} @{syntax goal_spec}? (@{syntax name} +) - ; - @@{method rotate_tac} @{syntax goal_spec}? @{syntax int}? - ; - (@@{method tactic} | @@{method raw_tactic}) @{syntax text} - ; - - dynamic_insts: ((@{syntax name} '=' @{syntax term}) + @'and') - "} - -\begin{description} - - \item @{method rule_tac} etc. do resolution of rules with explicit - instantiation. This works the same way as the ML tactics @{ML - res_inst_tac} etc. (see \cite{isabelle-implementation}) - - Multiple rules may be only given if there is no instantiation; then - @{method rule_tac} is the same as @{ML resolve_tac} in ML (see - \cite{isabelle-implementation}). - - \item @{method cut_tac} inserts facts into the proof state as - assumption of a subgoal; instantiations may be given as well. Note - that the scope of schematic variables is spread over the main goal - statement and rule premises are turned into new subgoals. This is - in contrast to the regular method @{method insert} which inserts - closed rule statements. - - \item @{method thin_tac}~@{text \} deletes the specified premise - from a subgoal. Note that @{text \} may contain schematic - variables, to abbreviate the intended proposition; the first - matching subgoal premise will be deleted. Removing useless premises - from a subgoal increases its readability and can make search tactics - run faster. - - \item @{method subgoal_tac}~@{text "\\<^sub>1 \ \\<^sub>n"} adds the propositions - @{text "\\<^sub>1 \ \\<^sub>n"} as local premises to a subgoal, and poses the same - as new subgoals (in the original context). - - \item @{method rename_tac}~@{text "x\<^sub>1 \ x\<^sub>n"} renames parameters of a - goal according to the list @{text "x\<^sub>1, \, x\<^sub>n"}, which refers to the - \emph{suffix} of variables. - - \item @{method rotate_tac}~@{text n} rotates the premises of a - subgoal by @{text n} positions: from right to left if @{text n} is - positive, and from left to right if @{text n} is negative; the - default value is 1. - - \item @{method tactic}~@{text "text"} produces a proof method from - any ML text of type @{ML_type tactic}. Apart from the usual ML - environment and the current proof context, the ML code may refer to - the locally bound values @{ML_text facts}, which indicates any - current facts used for forward-chaining. - - \item @{method raw_tactic} is similar to @{method tactic}, but - presents the goal state in its raw internal form, where simultaneous - subgoals appear as conjunction of the logical framework instead of - the usual split into several subgoals. While feature this is useful - for debugging of complex method definitions, it should not never - appear in production theories. - - \end{description} -*} - - -section {* The Simplifier \label{sec:simplifier} *} - -subsection {* Simplification methods *} - -text {* - \begin{matharray}{rcl} - @{method_def simp} & : & @{text method} \\ - @{method_def simp_all} & : & @{text method} \\ - \end{matharray} - - @{rail " - (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) - ; - - opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' - ; - @{syntax_def simpmod}: ('add' | 'del' | 'only' | 'cong' (() | 'add' | 'del') | - 'split' (() | 'add' | 'del')) ':' @{syntax thmrefs} - "} - - \begin{description} - - \item @{method simp} invokes the Simplifier, after declaring - additional rules according to the arguments given. Note that the - @{text only} modifier first removes all other rewrite rules, - congruences, and looper tactics (including splits), and then behaves - like @{text add}. - - \medskip The @{text cong} modifiers add or delete Simplifier - congruence rules (see also \secref{sec:simp-cong}), the default is - to add. - - \medskip The @{text split} modifiers add or delete rules for the - Splitter (see also \cite{isabelle-ref}), the default is to add. - This works only if the Simplifier method has been properly setup to - include the Splitter (all major object logics such HOL, HOLCF, FOL, - ZF do this already). - - \item @{method simp_all} is similar to @{method simp}, but acts on - all goals (backwards from the last to the first one). - - \end{description} - - By default the Simplifier methods take local assumptions fully into - account, using equational assumptions in the subsequent - normalization process, or simplifying assumptions themselves (cf.\ - @{ML asm_full_simp_tac} in \cite{isabelle-ref}). In structured - proofs this is usually quite well behaved in practice: just the - local premises of the actual goal are involved, additional facts may - be inserted via explicit forward-chaining (via @{command "then"}, - @{command "from"}, @{command "using"} etc.). - - Additional Simplifier options may be specified to tune the behavior - further (mostly for unstructured scripts with many accidental local - facts): ``@{text "(no_asm)"}'' means assumptions are ignored - completely (cf.\ @{ML simp_tac}), ``@{text "(no_asm_simp)"}'' means - assumptions are used in the simplification of the conclusion but are - not themselves simplified (cf.\ @{ML asm_simp_tac}), and ``@{text - "(no_asm_use)"}'' means assumptions are simplified but are not used - in the simplification of each other or the conclusion (cf.\ @{ML - full_simp_tac}). For compatibility reasons, there is also an option - ``@{text "(asm_lr)"}'', which means that an assumption is only used - for simplifying assumptions which are to the right of it (cf.\ @{ML - asm_lr_simp_tac}). - - The configuration option @{text "depth_limit"} limits the number of - recursive invocations of the simplifier during conditional - rewriting. - - \medskip The Splitter package is usually configured to work as part - of the Simplifier. The effect of repeatedly applying @{ML - split_tac} can be simulated by ``@{text "(simp only: split: - a\<^sub>1 \ a\<^sub>n)"}''. There is also a separate @{text split} - method available for single-step case splitting. -*} - - -subsection {* Declaring rules *} - -text {* - \begin{matharray}{rcl} - @{command_def "print_simpset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def simp} & : & @{text attribute} \\ - @{attribute_def split} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{attribute simp} | @@{attribute split}) (() | 'add' | 'del') - "} - - \begin{description} - - \item @{command "print_simpset"} prints the collection of rules - declared to the Simplifier, which is also known as ``simpset'' - internally \cite{isabelle-ref}. - - \item @{attribute simp} declares simplification rules. - - \item @{attribute split} declares case split rules. - - \end{description} -*} - - -subsection {* Congruence rules\label{sec:simp-cong} *} - -text {* - \begin{matharray}{rcl} - @{attribute_def cong} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{attribute cong} (() | 'add' | 'del') - "} - - \begin{description} - - \item @{attribute cong} declares congruence rules to the Simplifier - context. - - \end{description} - - Congruence rules are equalities of the form @{text [display] - "\ \ f ?x\<^sub>1 \ ?x\<^sub>n = f ?y\<^sub>1 \ ?y\<^sub>n"} - - This controls the simplification of the arguments of @{text f}. For - example, some arguments can be simplified under additional - assumptions: @{text [display] "?P\<^sub>1 \ ?Q\<^sub>1 \ (?Q\<^sub>1 \ ?P\<^sub>2 \ ?Q\<^sub>2) \ - (?P\<^sub>1 \ ?P\<^sub>2) \ (?Q\<^sub>1 \ ?Q\<^sub>2)"} - - Given this rule, the simplifier assumes @{text "?Q\<^sub>1"} and extracts - rewrite rules from it when simplifying @{text "?P\<^sub>2"}. Such local - assumptions are effective for rewriting formulae such as @{text "x = - 0 \ y + x = y"}. - - %FIXME - %The local assumptions are also provided as theorems to the solver; - %see \secref{sec:simp-solver} below. - - \medskip The following congruence rule for bounded quantifiers also - supplies contextual information --- about the bound variable: - @{text [display] "(?A = ?B) \ (\x. x \ ?B \ ?P x \ ?Q x) \ - (\x \ ?A. ?P x) \ (\x \ ?B. ?Q x)"} - - \medskip This congruence rule for conditional expressions can - supply contextual information for simplifying the arms: - @{text [display] "?p = ?q \ (?q \ ?a = ?c) \ (\ ?q \ ?b = ?d) \ - (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} - - A congruence rule can also \emph{prevent} simplification of some - arguments. Here is an alternative congruence rule for conditional - expressions that conforms to non-strict functional evaluation: - @{text [display] "?p = ?q \ (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} - - Only the first argument is simplified; the others remain unchanged. - This can make simplification much faster, but may require an extra - case split over the condition @{text "?q"} to prove the goal. -*} - - -subsection {* Simplification procedures *} - -text {* Simplification procedures are ML functions that produce proven - rewrite rules on demand. They are associated with higher-order - patterns that approximate the left-hand sides of equations. The - Simplifier first matches the current redex against one of the LHS - patterns; if this succeeds, the corresponding ML function is - invoked, passing the Simplifier context and redex term. Thus rules - may be specifically fashioned for particular situations, resulting - in a more powerful mechanism than term rewriting by a fixed set of - rules. - - Any successful result needs to be a (possibly conditional) rewrite - rule @{text "t \ u"} that is applicable to the current redex. The - rule will be applied just as any ordinary rewrite rule. It is - expected to be already in \emph{internal form}, bypassing the - automatic preprocessing of object-level equivalences. - - \begin{matharray}{rcl} - @{command_def "simproc_setup"} & : & @{text "local_theory \ local_theory"} \\ - simproc & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' - @{syntax text} \\ (@'identifier' (@{syntax nameref}+))? - ; - - @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) - "} - - \begin{description} - - \item @{command "simproc_setup"} defines a named simplification - procedure that is invoked by the Simplifier whenever any of the - given term patterns match the current redex. The implementation, - which is provided as ML source text, needs to be of type @{ML_type - "morphism -> simpset -> cterm -> thm option"}, where the @{ML_type - cterm} represents the current redex @{text r} and the result is - supposed to be some proven rewrite rule @{text "r \ r'"} (or a - generalized version), or @{ML NONE} to indicate failure. The - @{ML_type simpset} argument holds the full context of the current - Simplifier invocation, including the actual Isar proof context. The - @{ML_type morphism} informs about the difference of the original - compilation context wrt.\ the one of the actual application later - on. The optional @{keyword "identifier"} specifies theorems that - represent the logical content of the abstract theory of this - simproc. - - Morphisms and identifiers are only relevant for simprocs that are - defined within a local target context, e.g.\ in a locale. - - \item @{text "simproc add: name"} and @{text "simproc del: name"} - add or delete named simprocs to the current Simplifier context. The - default is to add a simproc. Note that @{command "simproc_setup"} - already adds the new simproc to the subsequent context. - - \end{description} -*} - - -subsubsection {* Example *} - -text {* The following simplification procedure for @{thm - [source=false, show_types] unit_eq} in HOL performs fine-grained - control over rule application, beyond higher-order pattern matching. - Declaring @{thm unit_eq} as @{attribute simp} directly would make - the simplifier loop! Note that a version of this simplification - procedure is already active in Isabelle/HOL. *} - -simproc_setup unit ("x::unit") = {* - fn _ => fn _ => fn ct => - if HOLogic.is_unit (term_of ct) then NONE - else SOME (mk_meta_eq @{thm unit_eq}) -*} - -text {* Since the Simplifier applies simplification procedures - frequently, it is important to make the failure check in ML - reasonably fast. *} - - -subsection {* Forward simplification *} - -text {* - \begin{matharray}{rcl} - @{attribute_def simplified} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{attribute simplified} opt? @{syntax thmrefs}? - ; - - opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' - "} - - \begin{description} - - \item @{attribute simplified}~@{text "a\<^sub>1 \ a\<^sub>n"} causes a theorem to - be simplified, either by exactly the specified rules @{text "a\<^sub>1, \, - a\<^sub>n"}, or the implicit Simplifier context if no arguments are given. - The result is fully simplified by default, including assumptions and - conclusion; the options @{text no_asm} etc.\ tune the Simplifier in - the same way as the for the @{text simp} method. - - Note that forward simplification restricts the simplifier to its - most basic operation of term rewriting; solver and looper tactics - \cite{isabelle-ref} are \emph{not} involved here. The @{text - simplified} attribute should be only rarely required under normal - circumstances. - - \end{description} -*} - - -section {* The Classical Reasoner \label{sec:classical} *} - -subsection {* Basic concepts *} - -text {* Although Isabelle is generic, many users will be working in - some extension of classical first-order logic. Isabelle/ZF is built - upon theory FOL, while Isabelle/HOL conceptually contains - first-order logic as a fragment. Theorem-proving in predicate logic - is undecidable, but many automated strategies have been developed to - assist in this task. - - Isabelle's classical reasoner is a generic package that accepts - certain information about a logic and delivers a suite of automatic - proof tools, based on rules that are classified and declared in the - context. These proof procedures are slow and simplistic compared - with high-end automated theorem provers, but they can save - considerable time and effort in practice. They can prove theorems - such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few - milliseconds (including full proof reconstruction): *} - -lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" - by blast - -lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" - by blast - -text {* The proof tools are generic. They are not restricted to - first-order logic, and have been heavily used in the development of - the Isabelle/HOL library and applications. The tactics can be - traced, and their components can be called directly; in this manner, - any proof can be viewed interactively. *} - - -subsubsection {* The sequent calculus *} - -text {* Isabelle supports natural deduction, which is easy to use for - interactive proof. But natural deduction does not easily lend - itself to automation, and has a bias towards intuitionism. For - certain proofs in classical logic, it can not be called natural. - The \emph{sequent calculus}, a generalization of natural deduction, - is easier to automate. - - A \textbf{sequent} has the form @{text "\ \ \"}, where @{text "\"} - and @{text "\"} are sets of formulae.\footnote{For first-order - logic, sequents can equivalently be made from lists or multisets of - formulae.} The sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} is - \textbf{valid} if @{text "P\<^sub>1 \ \ \ P\<^sub>m"} implies @{text "Q\<^sub>1 \ \ \ - Q\<^sub>n"}. Thus @{text "P\<^sub>1, \, P\<^sub>m"} represent assumptions, each of which - is true, while @{text "Q\<^sub>1, \, Q\<^sub>n"} represent alternative goals. A - sequent is \textbf{basic} if its left and right sides have a common - formula, as in @{text "P, Q \ Q, R"}; basic sequents are trivially - valid. - - Sequent rules are classified as \textbf{right} or \textbf{left}, - indicating which side of the @{text "\"} symbol they operate on. - Rules that operate on the right side are analogous to natural - deduction's introduction rules, and left rules are analogous to - elimination rules. The sequent calculus analogue of @{text "(\I)"} - is the rule - \[ - \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "P, \ \ \, Q"}} - \] - Applying the rule backwards, this breaks down some implication on - the right side of a sequent; @{text "\"} and @{text "\"} stand for - the sets of formulae that are unaffected by the inference. The - analogue of the pair @{text "(\I1)"} and @{text "(\I2)"} is the - single rule - \[ - \infer[@{text "(\R)"}]{@{text "\ \ \, P \ Q"}}{@{text "\ \ \, P, Q"}} - \] - This breaks down some disjunction on the right side, replacing it by - both disjuncts. Thus, the sequent calculus is a kind of - multiple-conclusion logic. - - To illustrate the use of multiple formulae on the right, let us - prove the classical theorem @{text "(P \ Q) \ (Q \ P)"}. Working - backwards, we reduce this formula to a basic sequent: - \[ - \infer[@{text "(\R)"}]{@{text "\ (P \ Q) \ (Q \ P)"}} - {\infer[@{text "(\R)"}]{@{text "\ (P \ Q), (Q \ P)"}} - {\infer[@{text "(\R)"}]{@{text "P \ Q, (Q \ P)"}} - {@{text "P, Q \ Q, P"}}}} - \] - - This example is typical of the sequent calculus: start with the - desired theorem and apply rules backwards in a fairly arbitrary - manner. This yields a surprisingly effective proof procedure. - Quantifiers add only few complications, since Isabelle handles - parameters and schematic variables. See \cite[Chapter - 10]{paulson-ml2} for further discussion. *} - - -subsubsection {* Simulating sequents by natural deduction *} - -text {* Isabelle can represent sequents directly, as in the - object-logic LK. But natural deduction is easier to work with, and - most object-logics employ it. Fortunately, we can simulate the - sequent @{text "P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n"} by the Isabelle formula - @{text "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1"} where the order of - the assumptions and the choice of @{text "Q\<^sub>1"} are arbitrary. - Elim-resolution plays a key role in simulating sequent proofs. - - We can easily handle reasoning on the left. Elim-resolution with - the rules @{text "(\E)"}, @{text "(\E)"} and @{text "(\E)"} achieves - a similar effect as the corresponding sequent rules. For the other - connectives, we use sequent-style elimination rules instead of - destruction rules such as @{text "(\E1, 2)"} and @{text "(\E)"}. - But note that the rule @{text "(\L)"} has no effect under our - representation of sequents! - \[ - \infer[@{text "(\L)"}]{@{text "\ P, \ \ \"}}{@{text "\ \ \, P"}} - \] - - What about reasoning on the right? Introduction rules can only - affect the formula in the conclusion, namely @{text "Q\<^sub>1"}. The - other right-side formulae are represented as negated assumptions, - @{text "\ Q\<^sub>2, \, \ Q\<^sub>n"}. In order to operate on one of these, it - must first be exchanged with @{text "Q\<^sub>1"}. Elim-resolution with the - @{text swap} rule has this effect: @{text "\ P \ (\ R \ P) \ R"} - - To ensure that swaps occur only when necessary, each introduction - rule is converted into a swapped form: it is resolved with the - second premise of @{text "(swap)"}. The swapped form of @{text - "(\I)"}, which might be called @{text "(\\E)"}, is - @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} - - Similarly, the swapped form of @{text "(\I)"} is - @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} - - Swapped introduction rules are applied using elim-resolution, which - deletes the negated formula. Our representation of sequents also - requires the use of ordinary introduction rules. If we had no - regard for readability of intermediate goal states, we could treat - the right side more uniformly by representing sequents as @{text - [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} -*} - - -subsubsection {* Extra rules for the sequent calculus *} - -text {* As mentioned, destruction rules such as @{text "(\E1, 2)"} and - @{text "(\E)"} must be replaced by sequent-style elimination rules. - In addition, we need rules to embody the classical equivalence - between @{text "P \ Q"} and @{text "\ P \ Q"}. The introduction - rules @{text "(\I1, 2)"} are replaced by a rule that simulates - @{text "(\R)"}: @{text [display] "(\ Q \ P) \ P \ Q"} - - The destruction rule @{text "(\E)"} is replaced by @{text [display] - "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} - - Quantifier replication also requires special rules. In classical - logic, @{text "\x. P x"} is equivalent to @{text "\ (\x. \ P x)"}; - the rules @{text "(\R)"} and @{text "(\L)"} are dual: - \[ - \infer[@{text "(\R)"}]{@{text "\ \ \, \x. P x"}}{@{text "\ \ \, \x. P x, P t"}} - \qquad - \infer[@{text "(\L)"}]{@{text "\x. P x, \ \ \"}}{@{text "P t, \x. P x, \ \ \"}} - \] - Thus both kinds of quantifier may be replicated. Theorems requiring - multiple uses of a universal formula are easy to invent; consider - @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any - @{text "n > 1"}. Natural examples of the multiple use of an - existential formula are rare; a standard one is @{text "\x. \y. P x - \ P y"}. - - Forgoing quantifier replication loses completeness, but gains - decidability, since the search space becomes finite. Many useful - theorems can be proved without replication, and the search generally - delivers its verdict in a reasonable time. To adopt this approach, - represent the sequent rules @{text "(\R)"}, @{text "(\L)"} and - @{text "(\R)"} by @{text "(\I)"}, @{text "(\E)"} and @{text "(\I)"}, - respectively, and put @{text "(\E)"} into elimination form: @{text - [display] "\x. P x \ (P t \ Q) \ Q"} - - Elim-resolution with this rule will delete the universal formula - after a single use. To replicate universal quantifiers, replace the - rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} - - To replicate existential quantifiers, replace @{text "(\I)"} by - @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} - - All introduction rules mentioned above are also useful in swapped - form. - - Replication makes the search space infinite; we must apply the rules - with care. The classical reasoner distinguishes between safe and - unsafe rules, applying the latter only when there is no alternative. - Depth-first search may well go down a blind alley; best-first search - is better behaved in an infinite search space. However, quantifier - replication is too expensive to prove any but the simplest theorems. -*} - - -subsection {* Rule declarations *} - -text {* The proof tools of the Classical Reasoner depend on - collections of rules declared in the context, which are classified - as introduction, elimination or destruction and as \emph{safe} or - \emph{unsafe}. In general, safe rules can be attempted blindly, - while unsafe rules must be used with care. A safe rule must never - reduce a provable goal to an unprovable set of subgoals. - - The rule @{text "P \ P \ Q"} is unsafe because it reduces @{text "P - \ Q"} to @{text "P"}, which might turn out as premature choice of an - unprovable subgoal. Any rule is unsafe whose premises contain new - unknowns. The elimination rule @{text "\x. P x \ (P t \ Q) \ Q"} is - unsafe, since it is applied via elim-resolution, which discards the - assumption @{text "\x. P x"} and replaces it by the weaker - assumption @{text "P t"}. The rule @{text "P t \ \x. P x"} is - unsafe for similar reasons. The quantifier duplication rule @{text - "\x. P x \ (P t \ \x. P x \ Q) \ Q"} is unsafe in a different sense: - since it keeps the assumption @{text "\x. P x"}, it is prone to - looping. In classical first-order logic, all rules are safe except - those mentioned above. - - The safe~/ unsafe distinction is vague, and may be regarded merely - as a way of giving some rules priority over others. One could argue - that @{text "(\E)"} is unsafe, because repeated application of it - could generate exponentially many subgoals. Induction rules are - unsafe because inductive proofs are difficult to set up - automatically. Any inference is unsafe that instantiates an unknown - in the proof state --- thus matching must be used, rather than - unification. Even proof by assumption is unsafe if it instantiates - unknowns shared with other subgoals. - - \begin{matharray}{rcl} - @{command_def "print_claset"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def intro} & : & @{text attribute} \\ - @{attribute_def elim} & : & @{text attribute} \\ - @{attribute_def dest} & : & @{text attribute} \\ - @{attribute_def rule} & : & @{text attribute} \\ - @{attribute_def iff} & : & @{text attribute} \\ - @{attribute_def swapped} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? - ; - @@{attribute rule} 'del' - ; - @@{attribute iff} (((() | 'add') '?'?) | 'del') - "} - - \begin{description} - - \item @{command "print_claset"} prints the collection of rules - declared to the Classical Reasoner, i.e.\ the @{ML_type claset} - within the context. - - \item @{attribute intro}, @{attribute elim}, and @{attribute dest} - declare introduction, elimination, and destruction rules, - respectively. By default, rules are considered as \emph{unsafe} - (i.e.\ not applied blindly without backtracking), while ``@{text - "!"}'' classifies as \emph{safe}. Rule declarations marked by - ``@{text "?"}'' coincide with those of Isabelle/Pure, cf.\ - \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps - of the @{method rule} method). The optional natural number - specifies an explicit weight argument, which is ignored by the - automated reasoning tools, but determines the search order of single - rule steps. - - Introduction rules are those that can be applied using ordinary - resolution. Their swapped forms are generated internally, which - will be applied using elim-resolution. Elimination rules are - applied using elim-resolution. Rules are sorted by the number of - new subgoals they will yield; rules that generate the fewest - subgoals will be tried first. Otherwise, later declarations take - precedence over earlier ones. - - Rules already present in the context with the same classification - are ignored. A warning is printed if the rule has already been - added with some other classification, but the rule is added anyway - as requested. - - \item @{attribute rule}~@{text del} deletes all occurrences of a - rule from the classical context, regardless of its classification as - introduction~/ elimination~/ destruction and safe~/ unsafe. - - \item @{attribute iff} declares logical equivalences to the - Simplifier and the Classical reasoner at the same time. - Non-conditional rules result in a safe introduction and elimination - pair; conditional ones are considered unsafe. Rules with negative - conclusion are automatically inverted (using @{text "\"}-elimination - internally). - - The ``@{text "?"}'' version of @{attribute iff} declares rules to - the Isabelle/Pure context only, and omits the Simplifier - declaration. - - \item @{attribute swapped} turns an introduction rule into an - elimination, by resolving with the classical swap principle @{text - "\ P \ (\ R \ P) \ R"} in the second position. This is mainly for - illustrative purposes: the Classical Reasoner already swaps rules - internally as explained above. - - \end{description} -*} - - -subsection {* Structured methods *} - -text {* - \begin{matharray}{rcl} - @{method_def rule} & : & @{text method} \\ - @{method_def contradiction} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method rule} @{syntax thmrefs}? - "} - - \begin{description} - - \item @{method rule} as offered by the Classical Reasoner is a - refinement over the Pure one (see \secref{sec:pure-meth-att}). Both - versions work the same, but the classical version observes the - classical rule context in addition to that of Isabelle/Pure. - - Common object logics (HOL, ZF, etc.) declare a rich collection of - classical rules (even if these would qualify as intuitionistic - ones), but only few declarations to the rule context of - Isabelle/Pure (\secref{sec:pure-meth-att}). - - \item @{method contradiction} solves some goal by contradiction, - deriving any result from both @{text "\ A"} and @{text A}. Chained - facts, which are guaranteed to participate, may appear in either - order. - - \end{description} -*} - - -subsection {* Automated methods *} - -text {* - \begin{matharray}{rcl} - @{method_def blast} & : & @{text method} \\ - @{method_def auto} & : & @{text method} \\ - @{method_def force} & : & @{text method} \\ - @{method_def fast} & : & @{text method} \\ - @{method_def slow} & : & @{text method} \\ - @{method_def best} & : & @{text method} \\ - @{method_def fastforce} & : & @{text method} \\ - @{method_def slowsimp} & : & @{text method} \\ - @{method_def bestsimp} & : & @{text method} \\ - @{method_def deepen} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method blast} @{syntax nat}? (@{syntax clamod} * ) - ; - @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) - ; - @@{method force} (@{syntax clasimpmod} * ) - ; - (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) - ; - (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) - (@{syntax clasimpmod} * ) - ; - @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) - ; - @{syntax_def clamod}: - (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thmrefs} - ; - @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | - ('cong' | 'split') (() | 'add' | 'del') | - 'iff' (((() | 'add') '?'?) | 'del') | - (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thmrefs} - "} - - \begin{description} - - \item @{method blast} is a separate classical tableau prover that - uses the same classical rule declarations as explained before. - - Proof search is coded directly in ML using special data structures. - A successful proof is then reconstructed using regular Isabelle - inferences. It is faster and more powerful than the other classical - reasoning tools, but has major limitations too. - - \begin{itemize} - - \item It does not use the classical wrapper tacticals, such as the - integration with the Simplifier of @{method fastforce}. - - \item It does not perform higher-order unification, as needed by the - rule @{thm [source=false] rangeI} in HOL. There are often - alternatives to such rules, for example @{thm [source=false] - range_eqI}. - - \item Function variables may only be applied to parameters of the - subgoal. (This restriction arises because the prover does not use - higher-order unification.) If other function variables are present - then the prover will fail with the message \texttt{Function Var's - argument not a bound variable}. - - \item Its proof strategy is more general than @{method fast} but can - be slower. If @{method blast} fails or seems to be running forever, - try @{method fast} and the other proof tools described below. - - \end{itemize} - - The optional integer argument specifies a bound for the number of - unsafe steps used in a proof. By default, @{method blast} starts - with a bound of 0 and increases it successively to 20. In contrast, - @{text "(blast lim)"} tries to prove the goal using a search bound - of @{text "lim"}. Sometimes a slow proof using @{method blast} can - be made much faster by supplying the successful search bound to this - proof method instead. - - \item @{method auto} combines classical reasoning with - simplification. It is intended for situations where there are a lot - of mostly trivial subgoals; it proves all the easy ones, leaving the - ones it cannot prove. Occasionally, attempting to prove the hard - ones may take a long time. - - The optional depth arguments in @{text "(auto m n)"} refer to its - builtin classical reasoning procedures: @{text m} (default 4) is for - @{method blast}, which is tried first, and @{text n} (default 2) is - for a slower but more general alternative that also takes wrappers - into account. - - \item @{method force} is intended to prove the first subgoal - completely, using many fancy proof tools and performing a rather - exhaustive search. As a result, proof attempts may take rather long - or diverge easily. - - \item @{method fast}, @{method best}, @{method slow} attempt to - prove the first subgoal using sequent-style reasoning as explained - before. Unlike @{method blast}, they construct proofs directly in - Isabelle. - - There is a difference in search strategy and back-tracking: @{method - fast} uses depth-first search and @{method best} uses best-first - search (guided by a heuristic function: normally the total size of - the proof state). - - Method @{method slow} is like @{method fast}, but conducts a broader - search: it may, when backtracking from a failed proof attempt, undo - even the step of proving a subgoal by assumption. - - \item @{method fastforce}, @{method slowsimp}, @{method bestsimp} - are like @{method fast}, @{method slow}, @{method best}, - respectively, but use the Simplifier as additional wrapper. The name - @{method fastforce}, reflects the behaviour of this popular method - better without requiring an understanding of its implementation. - - \item @{method deepen} works by exhaustive search up to a certain - depth. The start depth is 4 (unless specified explicitly), and the - depth is increased iteratively up to 10. Unsafe rules are modified - to preserve the formula they act on, so that it be used repeatedly. - This method can prove more goals than @{method fast}, but is much - slower, for example if the assumptions have many universal - quantifiers. - - \end{description} - - Any of the above methods support additional modifiers of the context - of classical (and simplifier) rules, but the ones related to the - Simplifier are explicitly prefixed by @{text simp} here. The - semantics of these ad-hoc rule declarations is analogous to the - attributes given before. Facts provided by forward chaining are - inserted into the goal before commencing proof search. -*} - - -subsection {* Semi-automated methods *} - -text {* These proof methods may help in situations when the - fully-automated tools fail. The result is a simpler subgoal that - can be tackled by other means, such as by manual instantiation of - quantifiers. - - \begin{matharray}{rcl} - @{method_def safe} & : & @{text method} \\ - @{method_def clarify} & : & @{text method} \\ - @{method_def clarsimp} & : & @{text method} \\ - \end{matharray} - - @{rail " - (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) - ; - @@{method clarsimp} (@{syntax clasimpmod} * ) - "} - - \begin{description} - - \item @{method safe} repeatedly performs safe steps on all subgoals. - It is deterministic, with at most one outcome. - - \item @{method clarify} performs a series of safe steps without - splitting subgoals; see also @{ML clarify_step_tac}. - - \item @{method clarsimp} acts like @{method clarify}, but also does - simplification. Note that if the Simplifier context includes a - splitter for the premises, the subgoal may still be split. - - \end{description} -*} - - -subsection {* Single-step tactics *} - -text {* - \begin{matharray}{rcl} - @{index_ML safe_step_tac: "Proof.context -> int -> tactic"} \\ - @{index_ML inst_step_tac: "Proof.context -> int -> tactic"} \\ - @{index_ML step_tac: "Proof.context -> int -> tactic"} \\ - @{index_ML slow_step_tac: "Proof.context -> int -> tactic"} \\ - @{index_ML clarify_step_tac: "Proof.context -> int -> tactic"} \\ - \end{matharray} - - These are the primitive tactics behind the (semi)automated proof - methods of the Classical Reasoner. By calling them yourself, you - can execute these procedures one step at a time. - - \begin{description} - - \item @{ML safe_step_tac}~@{text "ctxt i"} performs a safe step on - subgoal @{text i}. The safe wrapper tacticals are applied to a - tactic that may include proof by assumption or Modus Ponens (taking - care not to instantiate unknowns), or substitution. - - \item @{ML inst_step_tac} is like @{ML safe_step_tac}, but allows - unknowns to be instantiated. - - \item @{ML step_tac}~@{text "ctxt i"} is the basic step of the proof - procedure. The unsafe wrapper tacticals are applied to a tactic - that tries @{ML safe_tac}, @{ML inst_step_tac}, or applies an unsafe - rule from the context. - - \item @{ML slow_step_tac} resembles @{ML step_tac}, but allows - backtracking between using safe rules with instantiation (@{ML - inst_step_tac}) and using unsafe rules. The resulting search space - is larger. - - \item @{ML clarify_step_tac}~@{text "ctxt i"} performs a safe step - on subgoal @{text i}. No splitting step is applied; for example, - the subgoal @{text "A \ B"} is left as a conjunction. Proof by - assumption, Modus Ponens, etc., may be performed provided they do - not instantiate unknowns. Assumptions of the form @{text "x = t"} - may be eliminated. The safe wrapper tactical is applied. - - \end{description} -*} - - -section {* Object-logic setup \label{sec:object-logic} *} - -text {* - \begin{matharray}{rcl} - @{command_def "judgment"} & : & @{text "theory \ theory"} \\ - @{method_def atomize} & : & @{text method} \\ - @{attribute_def atomize} & : & @{text attribute} \\ - @{attribute_def rule_format} & : & @{text attribute} \\ - @{attribute_def rulify} & : & @{text attribute} \\ - \end{matharray} - - The very starting point for any Isabelle object-logic is a ``truth - judgment'' that links object-level statements to the meta-logic - (with its minimal language of @{text prop} that covers universal - quantification @{text "\"} and implication @{text "\"}). - - Common object-logics are sufficiently expressive to internalize rule - statements over @{text "\"} and @{text "\"} within their own - language. This is useful in certain situations where a rule needs - to be viewed as an atomic statement from the meta-level perspective, - e.g.\ @{text "\x. x \ A \ P x"} versus @{text "\x \ A. P x"}. - - From the following language elements, only the @{method atomize} - method and @{attribute rule_format} attribute are occasionally - required by end-users, the rest is for those who need to setup their - own object-logic. In the latter case existing formulations of - Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. - - Generic tools may refer to the information provided by object-logic - declarations internally. - - @{rail " - @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? - ; - @@{attribute atomize} ('(' 'full' ')')? - ; - @@{attribute rule_format} ('(' 'noasm' ')')? - "} - - \begin{description} - - \item @{command "judgment"}~@{text "c :: \ (mx)"} declares constant - @{text c} as the truth judgment of the current object-logic. Its - type @{text \} should specify a coercion of the category of - object-level propositions to @{text prop} of the Pure meta-logic; - the mixfix annotation @{text "(mx)"} would typically just link the - object language (internally of syntactic category @{text logic}) - with that of @{text prop}. Only one @{command "judgment"} - declaration may be given in any theory development. - - \item @{method atomize} (as a method) rewrites any non-atomic - premises of a sub-goal, using the meta-level equations declared via - @{attribute atomize} (as an attribute) beforehand. As a result, - heavily nested goals become amenable to fundamental operations such - as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``@{text - "(full)"}'' option here means to turn the whole subgoal into an - object-statement (if possible), including the outermost parameters - and assumptions as well. - - A typical collection of @{attribute atomize} rules for a particular - object-logic would provide an internalization for each of the - connectives of @{text "\"}, @{text "\"}, and @{text "\"}. - Meta-level conjunction should be covered as well (this is - particularly important for locales, see \secref{sec:locale}). - - \item @{attribute rule_format} rewrites a theorem by the equalities - declared as @{attribute rulify} rules in the current object-logic. - By default, the result is fully normalized, including assumptions - and conclusions at any depth. The @{text "(no_asm)"} option - restricts the transformation to the conclusion of a rule. - - In common object-logics (HOL, FOL, ZF), the effect of @{attribute - rule_format} is to replace (bounded) universal quantification - (@{text "\"}) and implication (@{text "\"}) by the corresponding - rule statements over @{text "\"} and @{text "\"}. - - \end{description} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/HOL_Specific.thy --- a/doc-src/IsarRef/Thy/HOL_Specific.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2366 +0,0 @@ -theory HOL_Specific -imports Base Main "~~/src/HOL/Library/Old_Recdef" -begin - -chapter {* Isabelle/HOL \label{ch:hol} *} - -section {* Higher-Order Logic *} - -text {* Isabelle/HOL is based on Higher-Order Logic, a polymorphic - version of Church's Simple Theory of Types. HOL can be best - understood as a simply-typed version of classical set theory. The - logic was first implemented in Gordon's HOL system - \cite{mgordon-hol}. It extends Church's original logic - \cite{church40} by explicit type variables (naive polymorphism) and - a sound axiomatization scheme for new types based on subsets of - existing types. - - Andrews's book \cite{andrews86} is a full description of the - original Church-style higher-order logic, with proofs of correctness - and completeness wrt.\ certain set-theoretic interpretations. The - particular extensions of Gordon-style HOL are explained semantically - in two chapters of the 1993 HOL book \cite{pitts93}. - - Experience with HOL over decades has demonstrated that higher-order - logic is widely applicable in many areas of mathematics and computer - science. In a sense, Higher-Order Logic is simpler than First-Order - Logic, because there are fewer restrictions and special cases. Note - that HOL is \emph{weaker} than FOL with axioms for ZF set theory, - which is traditionally considered the standard foundation of regular - mathematics, but for most applications this does not matter. If you - prefer ML to Lisp, you will probably prefer HOL to ZF. - - \medskip The syntax of HOL follows @{text "\"}-calculus and - functional programming. Function application is curried. To apply - the function @{text f} of type @{text "\\<^sub>1 \ \\<^sub>2 \ \\<^sub>3"} to the - arguments @{text a} and @{text b} in HOL, you simply write @{text "f - a b"} (as in ML or Haskell). There is no ``apply'' operator; the - existing application of the Pure @{text "\"}-calculus is re-used. - Note that in HOL @{text "f (a, b)"} means ``@{text "f"} applied to - the pair @{text "(a, b)"} (which is notation for @{text "Pair a - b"}). The latter typically introduces extra formal efforts that can - be avoided by currying functions by default. Explicit tuples are as - infrequent in HOL formalizations as in good ML or Haskell programs. - - \medskip Isabelle/HOL has a distinct feel, compared to other - object-logics like Isabelle/ZF. It identifies object-level types - with meta-level types, taking advantage of the default - type-inference mechanism of Isabelle/Pure. HOL fully identifies - object-level functions with meta-level functions, with native - abstraction and application. - - These identifications allow Isabelle to support HOL particularly - nicely, but they also mean that HOL requires some sophistication - from the user. In particular, an understanding of Hindley-Milner - type-inference with type-classes, which are both used extensively in - the standard libraries and applications. Beginners can set - @{attribute show_types} or even @{attribute show_sorts} to get more - explicit information about the result of type-inference. *} - - -section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "inductive"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "inductive_set"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "coinductive"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "coinductive_set"} & : & @{text "local_theory \ local_theory"} \\ - @{attribute_def (HOL) mono} & : & @{text attribute} \\ - \end{matharray} - - An \emph{inductive definition} specifies the least predicate or set - @{text R} closed under given rules: applying a rule to elements of - @{text R} yields a result within @{text R}. For example, a - structural operational semantics is an inductive definition of an - evaluation relation. - - Dually, a \emph{coinductive definition} specifies the greatest - predicate or set @{text R} that is consistent with given rules: - every element of @{text R} can be seen as arising by applying a rule - to elements of @{text R}. An important example is using - bisimulation relations to formalise equivalence of processes and - infinite data structures. - - Both inductive and coinductive definitions are based on the - Knaster-Tarski fixed-point theorem for complete lattices. The - collection of introduction rules given by the user determines a - functor on subsets of set-theoretic relations. The required - monotonicity of the recursion scheme is proven as a prerequisite to - the fixed-point definition and the resulting consequences. This - works by pushing inclusion through logical connectives and any other - operator that might be wrapped around recursive occurrences of the - defined relation: there must be a monotonicity theorem of the form - @{text "A \ B \ \ A \ \ B"}, for each premise @{text "\ R t"} in an - introduction rule. The default rule declarations of Isabelle/HOL - already take care of most common situations. - - @{rail " - (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | - @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) - @{syntax target}? \\ - @{syntax \"fixes\"} (@'for' @{syntax \"fixes\"})? (@'where' clauses)? \\ - (@'monos' @{syntax thmrefs})? - ; - clauses: (@{syntax thmdecl}? @{syntax prop} + '|') - ; - @@{attribute (HOL) mono} (() | 'add' | 'del') - "} - - \begin{description} - - \item @{command (HOL) "inductive"} and @{command (HOL) - "coinductive"} define (co)inductive predicates from the introduction - rules. - - The propositions given as @{text "clauses"} in the @{keyword - "where"} part are either rules of the usual @{text "\/\"} format - (with arbitrary nesting), or equalities using @{text "\"}. The - latter specifies extra-logical abbreviations in the sense of - @{command_ref abbreviation}. Introducing abstract syntax - simultaneously with the actual introduction rules is occasionally - useful for complex specifications. - - The optional @{keyword "for"} part contains a list of parameters of - the (co)inductive predicates that remain fixed throughout the - definition, in contrast to arguments of the relation that may vary - in each occurrence within the given @{text "clauses"}. - - The optional @{keyword "monos"} declaration contains additional - \emph{monotonicity theorems}, which are required for each operator - applied to a recursive set in the introduction rules. - - \item @{command (HOL) "inductive_set"} and @{command (HOL) - "coinductive_set"} are wrappers for to the previous commands for - native HOL predicates. This allows to define (co)inductive sets, - where multiple arguments are simulated via tuples. - - \item @{attribute (HOL) mono} declares monotonicity rules in the - context. These rule are involved in the automated monotonicity - proof of the above inductive and coinductive definitions. - - \end{description} -*} - - -subsection {* Derived rules *} - -text {* A (co)inductive definition of @{text R} provides the following - main theorems: - - \begin{description} - - \item @{text R.intros} is the list of introduction rules as proven - theorems, for the recursive predicates (or sets). The rules are - also available individually, using the names given them in the - theory file; - - \item @{text R.cases} is the case analysis (or elimination) rule; - - \item @{text R.induct} or @{text R.coinduct} is the (co)induction - rule; - - \item @{text R.simps} is the equation unrolling the fixpoint of the - predicate one step. - - \end{description} - - When several predicates @{text "R\<^sub>1, \, R\<^sub>n"} are - defined simultaneously, the list of introduction rules is called - @{text "R\<^sub>1_\_R\<^sub>n.intros"}, the case analysis rules are - called @{text "R\<^sub>1.cases, \, R\<^sub>n.cases"}, and the list - of mutual induction rules is called @{text - "R\<^sub>1_\_R\<^sub>n.inducts"}. -*} - - -subsection {* Monotonicity theorems *} - -text {* The context maintains a default set of theorems that are used - in monotonicity proofs. New rules can be declared via the - @{attribute (HOL) mono} attribute. See the main Isabelle/HOL - sources for some examples. The general format of such monotonicity - theorems is as follows: - - \begin{itemize} - - \item Theorems of the form @{text "A \ B \ \ A \ \ B"}, for proving - monotonicity of inductive definitions whose introduction rules have - premises involving terms such as @{text "\ R t"}. - - \item Monotonicity theorems for logical operators, which are of the - general form @{text "(\ \ \) \ \ (\ \ \) \ \ \ \"}. For example, in - the case of the operator @{text "\"}, the corresponding theorem is - \[ - \infer{@{text "P\<^sub>1 \ P\<^sub>2 \ Q\<^sub>1 \ Q\<^sub>2"}}{@{text "P\<^sub>1 \ Q\<^sub>1"} & @{text "P\<^sub>2 \ Q\<^sub>2"}} - \] - - \item De Morgan style equations for reasoning about the ``polarity'' - of expressions, e.g. - \[ - @{prop "\ \ P \ P"} \qquad\qquad - @{prop "\ (P \ Q) \ \ P \ \ Q"} - \] - - \item Equations for reducing complex operators to more primitive - ones whose monotonicity can easily be proved, e.g. - \[ - @{prop "(P \ Q) \ \ P \ Q"} \qquad\qquad - @{prop "Ball A P \ \x. x \ A \ P x"} - \] - - \end{itemize} -*} - -subsubsection {* Examples *} - -text {* The finite powerset operator can be defined inductively like this: *} - -inductive_set Fin :: "'a set \ 'a set set" for A :: "'a set" -where - empty: "{} \ Fin A" -| insert: "a \ A \ B \ Fin A \ insert a B \ Fin A" - -text {* The accessible part of a relation is defined as follows: *} - -inductive acc :: "('a \ 'a \ bool) \ 'a \ bool" - for r :: "'a \ 'a \ bool" (infix "\" 50) -where acc: "(\y. y \ x \ acc r y) \ acc r x" - -text {* Common logical connectives can be easily characterized as -non-recursive inductive definitions with parameters, but without -arguments. *} - -inductive AND for A B :: bool -where "A \ B \ AND A B" - -inductive OR for A B :: bool -where "A \ OR A B" - | "B \ OR A B" - -inductive EXISTS for B :: "'a \ bool" -where "B a \ EXISTS B" - -text {* Here the @{text "cases"} or @{text "induct"} rules produced by - the @{command inductive} package coincide with the expected - elimination rules for Natural Deduction. Already in the original - article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that - each connective can be characterized by its introductions, and the - elimination can be constructed systematically. *} - - -section {* Recursive functions \label{sec:recursion} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "primrec"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "fun"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def (HOL) "function"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def (HOL) "termination"} & : & @{text "local_theory \ proof(prove)"} \\ - \end{matharray} - - @{rail " - @@{command (HOL) primrec} @{syntax target}? @{syntax \"fixes\"} @'where' equations - ; - (@@{command (HOL) fun} | @@{command (HOL) function}) @{syntax target}? functionopts? - @{syntax \"fixes\"} \\ @'where' equations - ; - - equations: (@{syntax thmdecl}? @{syntax prop} + '|') - ; - functionopts: '(' (('sequential' | 'domintros') + ',') ')' - ; - @@{command (HOL) termination} @{syntax term}? - "} - - \begin{description} - - \item @{command (HOL) "primrec"} defines primitive recursive - functions over datatypes (see also @{command_ref (HOL) datatype} and - @{command_ref (HOL) rep_datatype}). The given @{text equations} - specify reduction rules that are produced by instantiating the - generic combinator for primitive recursion that is available for - each datatype. - - Each equation needs to be of the form: - - @{text [display] "f x\<^sub>1 \ x\<^sub>m (C y\<^sub>1 \ y\<^sub>k) z\<^sub>1 \ z\<^sub>n = rhs"} - - such that @{text C} is a datatype constructor, @{text rhs} contains - only the free variables on the left-hand side (or from the context), - and all recursive occurrences of @{text "f"} in @{text "rhs"} are of - the form @{text "f \ y\<^sub>i \"} for some @{text i}. At most one - reduction rule for each constructor can be given. The order does - not matter. For missing constructors, the function is defined to - return a default value, but this equation is made difficult to - access for users. - - The reduction rules are declared as @{attribute simp} by default, - which enables standard proof methods like @{method simp} and - @{method auto} to normalize expressions of @{text "f"} applied to - datatype constructions, by simulating symbolic computation via - rewriting. - - \item @{command (HOL) "function"} defines functions by general - wellfounded recursion. A detailed description with examples can be - found in \cite{isabelle-function}. The function is specified by a - set of (possibly conditional) recursive equations with arbitrary - pattern matching. The command generates proof obligations for the - completeness and the compatibility of patterns. - - The defined function is considered partial, and the resulting - simplification rules (named @{text "f.psimps"}) and induction rule - (named @{text "f.pinduct"}) are guarded by a generated domain - predicate @{text "f_dom"}. The @{command (HOL) "termination"} - command can then be used to establish that the function is total. - - \item @{command (HOL) "fun"} is a shorthand notation for ``@{command - (HOL) "function"}~@{text "(sequential)"}, followed by automated - proof attempts regarding pattern matching and termination. See - \cite{isabelle-function} for further details. - - \item @{command (HOL) "termination"}~@{text f} commences a - termination proof for the previously defined function @{text f}. If - this is omitted, the command refers to the most recent function - definition. After the proof is closed, the recursive equations and - the induction principle is established. - - \end{description} - - Recursive definitions introduced by the @{command (HOL) "function"} - command accommodate reasoning by induction (cf.\ @{method induct}): - rule @{text "f.induct"} refers to a specific induction rule, with - parameters named according to the user-specified equations. Cases - are numbered starting from 1. For @{command (HOL) "primrec"}, the - induction principle coincides with structural recursion on the - datatype where the recursion is carried out. - - The equations provided by these packages may be referred later as - theorem list @{text "f.simps"}, where @{text f} is the (collective) - name of the functions defined. Individual equations may be named - explicitly as well. - - The @{command (HOL) "function"} command accepts the following - options. - - \begin{description} - - \item @{text sequential} enables a preprocessor which disambiguates - overlapping patterns by making them mutually disjoint. Earlier - equations take precedence over later ones. This allows to give the - specification in a format very similar to functional programming. - Note that the resulting simplification and induction rules - correspond to the transformed specification, not the one given - originally. This usually means that each equation given by the user - may result in several theorems. Also note that this automatic - transformation only works for ML-style datatype patterns. - - \item @{text domintros} enables the automated generation of - introduction rules for the domain predicate. While mostly not - needed, they can be helpful in some proofs about partial functions. - - \end{description} -*} - -subsubsection {* Example: evaluation of expressions *} - -text {* Subsequently, we define mutual datatypes for arithmetic and - boolean expressions, and use @{command primrec} for evaluation - functions that follow the same recursive structure. *} - -datatype 'a aexp = - IF "'a bexp" "'a aexp" "'a aexp" - | Sum "'a aexp" "'a aexp" - | Diff "'a aexp" "'a aexp" - | Var 'a - | Num nat -and 'a bexp = - Less "'a aexp" "'a aexp" - | And "'a bexp" "'a bexp" - | Neg "'a bexp" - - -text {* \medskip Evaluation of arithmetic and boolean expressions *} - -primrec evala :: "('a \ nat) \ 'a aexp \ nat" - and evalb :: "('a \ nat) \ 'a bexp \ bool" -where - "evala env (IF b a1 a2) = (if evalb env b then evala env a1 else evala env a2)" -| "evala env (Sum a1 a2) = evala env a1 + evala env a2" -| "evala env (Diff a1 a2) = evala env a1 - evala env a2" -| "evala env (Var v) = env v" -| "evala env (Num n) = n" -| "evalb env (Less a1 a2) = (evala env a1 < evala env a2)" -| "evalb env (And b1 b2) = (evalb env b1 \ evalb env b2)" -| "evalb env (Neg b) = (\ evalb env b)" - -text {* Since the value of an expression depends on the value of its - variables, the functions @{const evala} and @{const evalb} take an - additional parameter, an \emph{environment} that maps variables to - their values. - - \medskip Substitution on expressions can be defined similarly. The - mapping @{text f} of type @{typ "'a \ 'a aexp"} given as a - parameter is lifted canonically on the types @{typ "'a aexp"} and - @{typ "'a bexp"}, respectively. -*} - -primrec substa :: "('a \ 'b aexp) \ 'a aexp \ 'b aexp" - and substb :: "('a \ 'b aexp) \ 'a bexp \ 'b bexp" -where - "substa f (IF b a1 a2) = IF (substb f b) (substa f a1) (substa f a2)" -| "substa f (Sum a1 a2) = Sum (substa f a1) (substa f a2)" -| "substa f (Diff a1 a2) = Diff (substa f a1) (substa f a2)" -| "substa f (Var v) = f v" -| "substa f (Num n) = Num n" -| "substb f (Less a1 a2) = Less (substa f a1) (substa f a2)" -| "substb f (And b1 b2) = And (substb f b1) (substb f b2)" -| "substb f (Neg b) = Neg (substb f b)" - -text {* In textbooks about semantics one often finds substitution - theorems, which express the relationship between substitution and - evaluation. For @{typ "'a aexp"} and @{typ "'a bexp"}, we can prove - such a theorem by mutual induction, followed by simplification. -*} - -lemma subst_one: - "evala env (substa (Var (v := a')) a) = evala (env (v := evala env a')) a" - "evalb env (substb (Var (v := a')) b) = evalb (env (v := evala env a')) b" - by (induct a and b) simp_all - -lemma subst_all: - "evala env (substa s a) = evala (\x. evala env (s x)) a" - "evalb env (substb s b) = evalb (\x. evala env (s x)) b" - by (induct a and b) simp_all - - -subsubsection {* Example: a substitution function for terms *} - -text {* Functions on datatypes with nested recursion are also defined - by mutual primitive recursion. *} - -datatype ('a, 'b) "term" = Var 'a | App 'b "('a, 'b) term list" - -text {* A substitution function on type @{typ "('a, 'b) term"} can be - defined as follows, by working simultaneously on @{typ "('a, 'b) - term list"}: *} - -primrec subst_term :: "('a \ ('a, 'b) term) \ ('a, 'b) term \ ('a, 'b) term" and - subst_term_list :: "('a \ ('a, 'b) term) \ ('a, 'b) term list \ ('a, 'b) term list" -where - "subst_term f (Var a) = f a" -| "subst_term f (App b ts) = App b (subst_term_list f ts)" -| "subst_term_list f [] = []" -| "subst_term_list f (t # ts) = subst_term f t # subst_term_list f ts" - -text {* The recursion scheme follows the structure of the unfolded - definition of type @{typ "('a, 'b) term"}. To prove properties of this - substitution function, mutual induction is needed: -*} - -lemma "subst_term (subst_term f1 \ f2) t = subst_term f1 (subst_term f2 t)" and - "subst_term_list (subst_term f1 \ f2) ts = subst_term_list f1 (subst_term_list f2 ts)" - by (induct t and ts) simp_all - - -subsubsection {* Example: a map function for infinitely branching trees *} - -text {* Defining functions on infinitely branching datatypes by - primitive recursion is just as easy. -*} - -datatype 'a tree = Atom 'a | Branch "nat \ 'a tree" - -primrec map_tree :: "('a \ 'b) \ 'a tree \ 'b tree" -where - "map_tree f (Atom a) = Atom (f a)" -| "map_tree f (Branch ts) = Branch (\x. map_tree f (ts x))" - -text {* Note that all occurrences of functions such as @{text ts} - above must be applied to an argument. In particular, @{term - "map_tree f \ ts"} is not allowed here. *} - -text {* Here is a simple composition lemma for @{term map_tree}: *} - -lemma "map_tree g (map_tree f t) = map_tree (g \ f) t" - by (induct t) simp_all - - -subsection {* Proof methods related to recursive definitions *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) pat_completeness} & : & @{text method} \\ - @{method_def (HOL) relation} & : & @{text method} \\ - @{method_def (HOL) lexicographic_order} & : & @{text method} \\ - @{method_def (HOL) size_change} & : & @{text method} \\ - @{method_def (HOL) induction_schema} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method (HOL) relation} @{syntax term} - ; - @@{method (HOL) lexicographic_order} (@{syntax clasimpmod} * ) - ; - @@{method (HOL) size_change} ( orders (@{syntax clasimpmod} * ) ) - ; - @@{method (HOL) induction_schema} - ; - orders: ( 'max' | 'min' | 'ms' ) * - "} - - \begin{description} - - \item @{method (HOL) pat_completeness} is a specialized method to - solve goals regarding the completeness of pattern matching, as - required by the @{command (HOL) "function"} package (cf.\ - \cite{isabelle-function}). - - \item @{method (HOL) relation}~@{text R} introduces a termination - proof using the relation @{text R}. The resulting proof state will - contain goals expressing that @{text R} is wellfounded, and that the - arguments of recursive calls decrease with respect to @{text R}. - Usually, this method is used as the initial proof step of manual - termination proofs. - - \item @{method (HOL) "lexicographic_order"} attempts a fully - automated termination proof by searching for a lexicographic - combination of size measures on the arguments of the function. The - method accepts the same arguments as the @{method auto} method, - which it uses internally to prove local descents. The @{syntax - clasimpmod} modifiers are accepted (as for @{method auto}). - - In case of failure, extensive information is printed, which can help - to analyse the situation (cf.\ \cite{isabelle-function}). - - \item @{method (HOL) "size_change"} also works on termination goals, - using a variation of the size-change principle, together with a - graph decomposition technique (see \cite{krauss_phd} for details). - Three kinds of orders are used internally: @{text max}, @{text min}, - and @{text ms} (multiset), which is only available when the theory - @{text Multiset} is loaded. When no order kinds are given, they are - tried in order. The search for a termination proof uses SAT solving - internally. - - For local descent proofs, the @{syntax clasimpmod} modifiers are - accepted (as for @{method auto}). - - \item @{method (HOL) induction_schema} derives user-specified - induction rules from well-founded induction and completeness of - patterns. This factors out some operations that are done internally - by the function package and makes them available separately. See - @{file "~~/src/HOL/ex/Induction_Schema.thy"} for examples. - - \end{description} -*} - - -subsection {* Functions with explicit partiality *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "partial_function"} & : & @{text "local_theory \ local_theory"} \\ - @{attribute_def (HOL) "partial_function_mono"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{command (HOL) partial_function} @{syntax target}? - '(' @{syntax nameref} ')' @{syntax \"fixes\"} \\ - @'where' @{syntax thmdecl}? @{syntax prop} - "} - - \begin{description} - - \item @{command (HOL) "partial_function"}~@{text "(mode)"} defines - recursive functions based on fixpoints in complete partial - orders. No termination proof is required from the user or - constructed internally. Instead, the possibility of non-termination - is modelled explicitly in the result type, which contains an - explicit bottom element. - - Pattern matching and mutual recursion are currently not supported. - Thus, the specification consists of a single function described by a - single recursive equation. - - There are no fixed syntactic restrictions on the body of the - function, but the induced functional must be provably monotonic - wrt.\ the underlying order. The monotonicitity proof is performed - internally, and the definition is rejected when it fails. The proof - can be influenced by declaring hints using the - @{attribute (HOL) partial_function_mono} attribute. - - The mandatory @{text mode} argument specifies the mode of operation - of the command, which directly corresponds to a complete partial - order on the result type. By default, the following modes are - defined: - - \begin{description} - - \item @{text option} defines functions that map into the @{type - option} type. Here, the value @{term None} is used to model a - non-terminating computation. Monotonicity requires that if @{term - None} is returned by a recursive call, then the overall result must - also be @{term None}. This is best achieved through the use of the - monadic operator @{const "Option.bind"}. - - \item @{text tailrec} defines functions with an arbitrary result - type and uses the slightly degenerated partial order where @{term - "undefined"} is the bottom element. Now, monotonicity requires that - if @{term undefined} is returned by a recursive call, then the - overall result must also be @{term undefined}. In practice, this is - only satisfied when each recursive call is a tail call, whose result - is directly returned. Thus, this mode of operation allows the - definition of arbitrary tail-recursive functions. - - \end{description} - - Experienced users may define new modes by instantiating the locale - @{const "partial_function_definitions"} appropriately. - - \item @{attribute (HOL) partial_function_mono} declares rules for - use in the internal monononicity proofs of partial function - definitions. - - \end{description} - -*} - - -subsection {* Old-style recursive function definitions (TFL) *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "recdef"} & : & @{text "theory \ theory)"} \\ - @{command_def (HOL) "recdef_tc"}@{text "\<^sup>*"} & : & @{text "theory \ proof(prove)"} \\ - \end{matharray} - - The old TFL commands @{command (HOL) "recdef"} and @{command (HOL) - "recdef_tc"} for defining recursive are mostly obsolete; @{command - (HOL) "function"} or @{command (HOL) "fun"} should be used instead. - - @{rail " - @@{command (HOL) recdef} ('(' @'permissive' ')')? \\ - @{syntax name} @{syntax term} (@{syntax prop} +) hints? - ; - recdeftc @{syntax thmdecl}? tc - ; - hints: '(' @'hints' ( recdefmod * ) ')' - ; - recdefmod: (('recdef_simp' | 'recdef_cong' | 'recdef_wf') - (() | 'add' | 'del') ':' @{syntax thmrefs}) | @{syntax clasimpmod} - ; - tc: @{syntax nameref} ('(' @{syntax nat} ')')? - "} - - \begin{description} - - \item @{command (HOL) "recdef"} defines general well-founded - recursive functions (using the TFL package), see also - \cite{isabelle-HOL}. The ``@{text "(permissive)"}'' option tells - TFL to recover from failed proof attempts, returning unfinished - results. The @{text recdef_simp}, @{text recdef_cong}, and @{text - recdef_wf} hints refer to auxiliary rules to be used in the internal - automated proof process of TFL. Additional @{syntax clasimpmod} - declarations may be given to tune the context of the Simplifier - (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ - \secref{sec:classical}). - - \item @{command (HOL) "recdef_tc"}~@{text "c (i)"} recommences the - proof for leftover termination condition number @{text i} (default - 1) as generated by a @{command (HOL) "recdef"} definition of - constant @{text c}. - - Note that in most cases, @{command (HOL) "recdef"} is able to finish - its internal proofs without manual intervention. - - \end{description} - - \medskip Hints for @{command (HOL) "recdef"} may be also declared - globally, using the following attributes. - - \begin{matharray}{rcl} - @{attribute_def (HOL) recdef_simp} & : & @{text attribute} \\ - @{attribute_def (HOL) recdef_cong} & : & @{text attribute} \\ - @{attribute_def (HOL) recdef_wf} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - (@@{attribute (HOL) recdef_simp} | @@{attribute (HOL) recdef_cong} | - @@{attribute (HOL) recdef_wf}) (() | 'add' | 'del') - "} -*} - - -section {* Datatypes \label{sec:hol-datatype} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "datatype"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "rep_datatype"} & : & @{text "theory \ proof(prove)"} \\ - \end{matharray} - - @{rail " - @@{command (HOL) datatype} (spec + @'and') - ; - @@{command (HOL) rep_datatype} ('(' (@{syntax name} +) ')')? (@{syntax term} +) - ; - - spec: @{syntax typespec_sorts} @{syntax mixfix}? '=' (cons + '|') - ; - cons: @{syntax name} (@{syntax type} * ) @{syntax mixfix}? - "} - - \begin{description} - - \item @{command (HOL) "datatype"} defines inductive datatypes in - HOL. - - \item @{command (HOL) "rep_datatype"} represents existing types as - datatypes. - - For foundational reasons, some basic types such as @{typ nat}, @{typ - "'a \ 'b"}, @{typ "'a + 'b"}, @{typ bool} and @{typ unit} are - introduced by more primitive means using @{command_ref typedef}. To - recover the rich infrastructure of @{command datatype} (e.g.\ rules - for @{method cases} and @{method induct} and the primitive recursion - combinators), such types may be represented as actual datatypes - later. This is done by specifying the constructors of the desired - type, and giving a proof of the induction rule, distinctness and - injectivity of constructors. - - For example, see @{file "~~/src/HOL/Sum_Type.thy"} for the - representation of the primitive sum type as fully-featured datatype. - - \end{description} - - The generated rules for @{method induct} and @{method cases} provide - case names according to the given constructors, while parameters are - named after the types (see also \secref{sec:cases-induct}). - - See \cite{isabelle-HOL} for more details on datatypes, but beware of - the old-style theory syntax being used there! Apart from proper - proof methods for case-analysis and induction, there are also - emulations of ML tactics @{method (HOL) case_tac} and @{method (HOL) - induct_tac} available, see \secref{sec:hol-induct-tac}; these admit - to refer directly to the internal structure of subgoals (including - internally bound parameters). -*} - - -subsubsection {* Examples *} - -text {* We define a type of finite sequences, with slightly different - names than the existing @{typ "'a list"} that is already in @{theory - Main}: *} - -datatype 'a seq = Empty | Seq 'a "'a seq" - -text {* We can now prove some simple lemma by structural induction: *} - -lemma "Seq x xs \ xs" -proof (induct xs arbitrary: x) - case Empty - txt {* This case can be proved using the simplifier: the freeness - properties of the datatype are already declared as @{attribute - simp} rules. *} - show "Seq x Empty \ Empty" - by simp -next - case (Seq y ys) - txt {* The step case is proved similarly. *} - show "Seq x (Seq y ys) \ Seq y ys" - using `Seq y ys \ ys` by simp -qed - -text {* Here is a more succinct version of the same proof: *} - -lemma "Seq x xs \ xs" - by (induct xs arbitrary: x) simp_all - - -section {* Records \label{sec:hol-record} *} - -text {* - In principle, records merely generalize the concept of tuples, where - components may be addressed by labels instead of just position. The - logical infrastructure of records in Isabelle/HOL is slightly more - advanced, though, supporting truly extensible record schemes. This - admits operations that are polymorphic with respect to record - extension, yielding ``object-oriented'' effects like (single) - inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more - details on object-oriented verification and record subtyping in HOL. -*} - - -subsection {* Basic concepts *} - -text {* - Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records - at the level of terms and types. The notation is as follows: - - \begin{center} - \begin{tabular}{l|l|l} - & record terms & record types \\ \hline - fixed & @{text "\x = a, y = b\"} & @{text "\x :: A, y :: B\"} \\ - schematic & @{text "\x = a, y = b, \ = m\"} & - @{text "\x :: A, y :: B, \ :: M\"} \\ - \end{tabular} - \end{center} - - \noindent The ASCII representation of @{text "\x = a\"} is @{text - "(| x = a |)"}. - - A fixed record @{text "\x = a, y = b\"} has field @{text x} of value - @{text a} and field @{text y} of value @{text b}. The corresponding - type is @{text "\x :: A, y :: B\"}, assuming that @{text "a :: A"} - and @{text "b :: B"}. - - A record scheme like @{text "\x = a, y = b, \ = m\"} contains fields - @{text x} and @{text y} as before, but also possibly further fields - as indicated by the ``@{text "\"}'' notation (which is actually part - of the syntax). The improper field ``@{text "\"}'' of a record - scheme is called the \emph{more part}. Logically it is just a free - variable, which is occasionally referred to as ``row variable'' in - the literature. The more part of a record scheme may be - instantiated by zero or more further components. For example, the - previous scheme may get instantiated to @{text "\x = a, y = b, z = - c, \ = m'\"}, where @{text m'} refers to a different more part. - Fixed records are special instances of record schemes, where - ``@{text "\"}'' is properly terminated by the @{text "() :: unit"} - element. In fact, @{text "\x = a, y = b\"} is just an abbreviation - for @{text "\x = a, y = b, \ = ()\"}. - - \medskip Two key observations make extensible records in a simply - typed language like HOL work out: - - \begin{enumerate} - - \item the more part is internalized, as a free term or type - variable, - - \item field names are externalized, they cannot be accessed within - the logic as first-class values. - - \end{enumerate} - - \medskip In Isabelle/HOL record types have to be defined explicitly, - fixing their field names and types, and their (optional) parent - record. Afterwards, records may be formed using above syntax, while - obeying the canonical order of fields as given by their declaration. - The record package provides several standard operations like - selectors and updates. The common setup for various generic proof - tools enable succinct reasoning patterns. See also the Isabelle/HOL - tutorial \cite{isabelle-hol-book} for further instructions on using - records in practice. -*} - - -subsection {* Record specifications *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "record"} & : & @{text "theory \ theory"} \\ - \end{matharray} - - @{rail " - @@{command (HOL) record} @{syntax typespec_sorts} '=' \\ - (@{syntax type} '+')? (constdecl +) - ; - constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? - "} - - \begin{description} - - \item @{command (HOL) "record"}~@{text "(\\<^sub>1, \, \\<^sub>m) t = \ + c\<^sub>1 :: \\<^sub>1 - \ c\<^sub>n :: \\<^sub>n"} defines extensible record type @{text "(\\<^sub>1, \, \\<^sub>m) t"}, - derived from the optional parent record @{text "\"} by adding new - field components @{text "c\<^sub>i :: \\<^sub>i"} etc. - - The type variables of @{text "\"} and @{text "\\<^sub>i"} need to be - covered by the (distinct) parameters @{text "\\<^sub>1, \, - \\<^sub>m"}. Type constructor @{text t} has to be new, while @{text - \} needs to specify an instance of an existing record type. At - least one new field @{text "c\<^sub>i"} has to be specified. - Basically, field names need to belong to a unique record. This is - not a real restriction in practice, since fields are qualified by - the record name internally. - - The parent record specification @{text \} is optional; if omitted - @{text t} becomes a root record. The hierarchy of all records - declared within a theory context forms a forest structure, i.e.\ a - set of trees starting with a root record each. There is no way to - merge multiple parent records! - - For convenience, @{text "(\\<^sub>1, \, \\<^sub>m) t"} is made a - type abbreviation for the fixed record type @{text "\c\<^sub>1 :: - \\<^sub>1, \, c\<^sub>n :: \\<^sub>n\"}, likewise is @{text - "(\\<^sub>1, \, \\<^sub>m, \) t_scheme"} made an abbreviation for - @{text "\c\<^sub>1 :: \\<^sub>1, \, c\<^sub>n :: \\<^sub>n, \ :: - \\"}. - - \end{description} -*} - - -subsection {* Record operations *} - -text {* - Any record definition of the form presented above produces certain - standard operations. Selectors and updates are provided for any - field, including the improper one ``@{text more}''. There are also - cumulative record constructor functions. To simplify the - presentation below, we assume for now that @{text "(\\<^sub>1, \, - \\<^sub>m) t"} is a root record with fields @{text "c\<^sub>1 :: - \\<^sub>1, \, c\<^sub>n :: \\<^sub>n"}. - - \medskip \textbf{Selectors} and \textbf{updates} are available for - any field (including ``@{text more}''): - - \begin{matharray}{lll} - @{text "c\<^sub>i"} & @{text "::"} & @{text "\\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i"} \\ - @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\\<^sub>i \ \\<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>c :: \<^vec>\, \ :: \\"} \\ - \end{matharray} - - There is special syntax for application of updates: @{text "r\x := - a\"} abbreviates term @{text "x_update a r"}. Further notation for - repeated updates is also available: @{text "r\x := a\\y := b\\z := - c\"} may be written @{text "r\x := a, y := b, z := c\"}. Note that - because of postfix notation the order of fields shown here is - reverse than in the actual term. Since repeated updates are just - function applications, fields may be freely permuted in @{text "\x - := a, y := b, z := c\"}, as far as logical equality is concerned. - Thus commutativity of independent updates can be proven within the - logic for any two fields, but not as a general theorem. - - \medskip The \textbf{make} operation provides a cumulative record - constructor function: - - \begin{matharray}{lll} - @{text "t.make"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\"} \\ - \end{matharray} - - \medskip We now reconsider the case of non-root records, which are - derived of some parent. In general, the latter may depend on - another parent as well, resulting in a list of \emph{ancestor - records}. Appending the lists of fields of all ancestors results in - a certain field prefix. The record package automatically takes care - of this by lifting operations over this context of ancestor fields. - Assuming that @{text "(\\<^sub>1, \, \\<^sub>m) t"} has ancestor - fields @{text "b\<^sub>1 :: \\<^sub>1, \, b\<^sub>k :: \\<^sub>k"}, - the above record operations will get the following types: - - \medskip - \begin{tabular}{lll} - @{text "c\<^sub>i"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^sub>i"} \\ - @{text "c\<^sub>i_update"} & @{text "::"} & @{text "\\<^sub>i \ - \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ - \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\"} \\ - @{text "t.make"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>k \ \\<^sub>1 \ \ \\<^sub>n \ - \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\"} \\ - \end{tabular} - \medskip - - \noindent Some further operations address the extension aspect of a - derived record scheme specifically: @{text "t.fields"} produces a - record fragment consisting of exactly the new fields introduced here - (the result may serve as a more part elsewhere); @{text "t.extend"} - takes a fixed record and adds a given more part; @{text - "t.truncate"} restricts a record scheme to a fixed record. - - \medskip - \begin{tabular}{lll} - @{text "t.fields"} & @{text "::"} & @{text "\\<^sub>1 \ \ \\<^sub>n \ \\<^vec>c :: \<^vec>\\"} \\ - @{text "t.extend"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\ \ - \ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\"} \\ - @{text "t.truncate"} & @{text "::"} & @{text "\\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\, \ :: \\ \ \\<^vec>b :: \<^vec>\, \<^vec>c :: \<^vec>\\"} \\ - \end{tabular} - \medskip - - \noindent Note that @{text "t.make"} and @{text "t.fields"} coincide - for root records. -*} - - -subsection {* Derived rules and proof tools *} - -text {* - The record package proves several results internally, declaring - these facts to appropriate proof tools. This enables users to - reason about record structures quite conveniently. Assume that - @{text t} is a record type as specified above. - - \begin{enumerate} - - \item Standard conversions for selectors or updates applied to - record constructor terms are made part of the default Simplifier - context; thus proofs by reduction of basic operations merely require - the @{method simp} method without further arguments. These rules - are available as @{text "t.simps"}, too. - - \item Selectors applied to updated records are automatically reduced - by an internal simplification procedure, which is also part of the - standard Simplifier setup. - - \item Inject equations of a form analogous to @{prop "(x, y) = (x', - y') \ x = x' \ y = y'"} are declared to the Simplifier and Classical - Reasoner as @{attribute iff} rules. These rules are available as - @{text "t.iffs"}. - - \item The introduction rule for record equality analogous to @{text - "x r = x r' \ y r = y r' \ \ r = r'"} is declared to the Simplifier, - and as the basic rule context as ``@{attribute intro}@{text "?"}''. - The rule is called @{text "t.equality"}. - - \item Representations of arbitrary record expressions as canonical - constructor terms are provided both in @{method cases} and @{method - induct} format (cf.\ the generic proof methods of the same name, - \secref{sec:cases-induct}). Several variations are available, for - fixed records, record schemes, more parts etc. - - The generic proof methods are sufficiently smart to pick the most - sensible rule according to the type of the indicated record - expression: users just need to apply something like ``@{text "(cases - r)"}'' to a certain proof problem. - - \item The derived record operations @{text "t.make"}, @{text - "t.fields"}, @{text "t.extend"}, @{text "t.truncate"} are \emph{not} - treated automatically, but usually need to be expanded by hand, - using the collective fact @{text "t.defs"}. - - \end{enumerate} -*} - - -subsubsection {* Examples *} - -text {* See @{file "~~/src/HOL/ex/Records.thy"}, for example. *} - - -section {* Adhoc tuples *} - -text {* - \begin{matharray}{rcl} - @{attribute_def (HOL) split_format}@{text "\<^sup>*"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{attribute (HOL) split_format} ('(' 'complete' ')')? - "} - - \begin{description} - - \item @{attribute (HOL) split_format}\ @{text "(complete)"} causes - arguments in function applications to be represented canonically - according to their tuple type structure. - - Note that this operation tends to invent funny names for new local - parameters introduced. - - \end{description} -*} - - -section {* Typedef axiomatization \label{sec:hol-typedef} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "typedef"} & : & @{text "local_theory \ proof(prove)"} \\ - \end{matharray} - - A Gordon/HOL-style type definition is a certain axiom scheme that - identifies a new type with a subset of an existing type. More - precisely, the new type is defined by exhibiting an existing type - @{text \}, a set @{text "A :: \ set"}, and a theorem that proves - @{prop "\x. x \ A"}. Thus @{text A} is a non-empty subset of @{text - \}, and the new type denotes this subset. New functions are - postulated that establish an isomorphism between the new type and - the subset. In general, the type @{text \} may involve type - variables @{text "\\<^sub>1, \, \\<^sub>n"} which means that the type definition - produces a type constructor @{text "(\\<^sub>1, \, \\<^sub>n) t"} depending on - those type arguments. - - The axiomatization can be considered a ``definition'' in the sense - of the particular set-theoretic interpretation of HOL - \cite{pitts93}, where the universe of types is required to be - downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely - new types introduced by @{command "typedef"} stay within the range - of HOL models by construction. Note that @{command_ref - type_synonym} from Isabelle/Pure merely introduces syntactic - abbreviations, without any logical significance. - - @{rail " - @@{command (HOL) typedef} alt_name? abs_type '=' rep_set - ; - - alt_name: '(' (@{syntax name} | @'open' | @'open' @{syntax name}) ')' - ; - abs_type: @{syntax typespec_sorts} @{syntax mixfix}? - ; - rep_set: @{syntax term} (@'morphisms' @{syntax name} @{syntax name})? - "} - - \begin{description} - - \item @{command (HOL) "typedef"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = A"} - axiomatizes a type definition in the background theory of the - current context, depending on a non-emptiness result of the set - @{text A} that needs to be proven here. The set @{text A} may - contain type variables @{text "\\<^sub>1, \, \\<^sub>n"} as specified on the LHS, - but no term variables. - - Even though a local theory specification, the newly introduced type - constructor cannot depend on parameters or assumptions of the - context: this is structurally impossible in HOL. In contrast, the - non-emptiness proof may use local assumptions in unusual situations, - which could result in different interpretations in target contexts: - the meaning of the bijection between the representing set @{text A} - and the new type @{text t} may then change in different application - contexts. - - By default, @{command (HOL) "typedef"} defines both a type - constructor @{text t} for the new type, and a term constant @{text - t} for the representing set within the old type. Use the ``@{text - "(open)"}'' option to suppress a separate constant definition - altogether. The injection from type to set is called @{text Rep_t}, - its inverse @{text Abs_t}, unless explicit @{keyword (HOL) - "morphisms"} specification provides alternative names. - - The core axiomatization uses the locale predicate @{const - type_definition} as defined in Isabelle/HOL. Various basic - consequences of that are instantiated accordingly, re-using the - locale facts with names derived from the new type constructor. Thus - the generic @{thm type_definition.Rep} is turned into the specific - @{text "Rep_t"}, for example. - - Theorems @{thm type_definition.Rep}, @{thm - type_definition.Rep_inverse}, and @{thm type_definition.Abs_inverse} - provide the most basic characterization as a corresponding - injection/surjection pair (in both directions). The derived rules - @{thm type_definition.Rep_inject} and @{thm - type_definition.Abs_inject} provide a more convenient version of - injectivity, suitable for automated proof tools (e.g.\ in - declarations involving @{attribute simp} or @{attribute iff}). - Furthermore, the rules @{thm type_definition.Rep_cases}~/ @{thm - type_definition.Rep_induct}, and @{thm type_definition.Abs_cases}~/ - @{thm type_definition.Abs_induct} provide alternative views on - surjectivity. These rules are already declared as set or type rules - for the generic @{method cases} and @{method induct} methods, - respectively. - - An alternative name for the set definition (and other derived - entities) may be specified in parentheses; the default is to use - @{text t} directly. - - \end{description} - - \begin{warn} - If you introduce a new type axiomatically, i.e.\ via @{command_ref - typedecl} and @{command_ref axiomatization}, the minimum requirement - is that it has a non-empty model, to avoid immediate collapse of the - HOL logic. Moreover, one needs to demonstrate that the - interpretation of such free-form axiomatizations can coexist with - that of the regular @{command_def typedef} scheme, and any extension - that other people might have introduced elsewhere (e.g.\ in HOLCF - \cite{MuellerNvOS99}). - \end{warn} -*} - -subsubsection {* Examples *} - -text {* Type definitions permit the introduction of abstract data - types in a safe way, namely by providing models based on already - existing types. Given some abstract axiomatic description @{text P} - of a type, this involves two steps: - - \begin{enumerate} - - \item Find an appropriate type @{text \} and subset @{text A} which - has the desired properties @{text P}, and make a type definition - based on this representation. - - \item Prove that @{text P} holds for @{text \} by lifting @{text P} - from the representation. - - \end{enumerate} - - You can later forget about the representation and work solely in - terms of the abstract properties @{text P}. - - \medskip The following trivial example pulls a three-element type - into existence within the formal logical environment of HOL. *} - -typedef three = "{(True, True), (True, False), (False, True)}" - by blast - -definition "One = Abs_three (True, True)" -definition "Two = Abs_three (True, False)" -definition "Three = Abs_three (False, True)" - -lemma three_distinct: "One \ Two" "One \ Three" "Two \ Three" - by (simp_all add: One_def Two_def Three_def Abs_three_inject three_def) - -lemma three_cases: - fixes x :: three obtains "x = One" | "x = Two" | "x = Three" - by (cases x) (auto simp: One_def Two_def Three_def Abs_three_inject three_def) - -text {* Note that such trivial constructions are better done with - derived specification mechanisms such as @{command datatype}: *} - -datatype three' = One' | Two' | Three' - -text {* This avoids re-doing basic definitions and proofs from the - primitive @{command typedef} above. *} - - -section {* Functorial structure of types *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "enriched_type"} & : & @{text "local_theory \ proof(prove)"} - \end{matharray} - - @{rail " - @@{command (HOL) enriched_type} (@{syntax name} ':')? @{syntax term} - ; - "} - - \begin{description} - - \item @{command (HOL) "enriched_type"}~@{text "prefix: m"} allows to - prove and register properties about the functorial structure of type - constructors. These properties then can be used by other packages - to deal with those type constructors in certain type constructions. - Characteristic theorems are noted in the current local theory. By - default, they are prefixed with the base name of the type - constructor, an explicit prefix can be given alternatively. - - The given term @{text "m"} is considered as \emph{mapper} for the - corresponding type constructor and must conform to the following - type pattern: - - \begin{matharray}{lll} - @{text "m"} & @{text "::"} & - @{text "\\<^isub>1 \ \ \\<^isub>k \ (\<^vec>\\<^isub>n) t \ (\<^vec>\\<^isub>n) t"} \\ - \end{matharray} - - \noindent where @{text t} is the type constructor, @{text - "\<^vec>\\<^isub>n"} and @{text "\<^vec>\\<^isub>n"} are distinct - type variables free in the local theory and @{text "\\<^isub>1"}, - \ldots, @{text "\\<^isub>k"} is a subsequence of @{text "\\<^isub>1 \ - \\<^isub>1"}, @{text "\\<^isub>1 \ \\<^isub>1"}, \ldots, - @{text "\\<^isub>n \ \\<^isub>n"}, @{text "\\<^isub>n \ - \\<^isub>n"}. - - \end{description} -*} - - -section {* Transfer package *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) "transfer"} & : & @{text method} \\ - @{method_def (HOL) "transfer'"} & : & @{text method} \\ - @{method_def (HOL) "transfer_prover"} & : & @{text method} \\ - @{attribute_def (HOL) "transfer_rule"} & : & @{text attribute} \\ - @{attribute_def (HOL) "relator_eq"} & : & @{text attribute} \\ - \end{matharray} - - \begin{description} - - \item @{method (HOL) "transfer"} method replaces the current subgoal - with a logically equivalent one that uses different types and - constants. The replacement of types and constants is guided by the - database of transfer rules. Goals are generalized over all free - variables by default; this is necessary for variables whose types - change, but can be overridden for specific variables with e.g. - @{text "transfer fixing: x y z"}. - - \item @{method (HOL) "transfer'"} is a variant of @{method (HOL) - transfer} that allows replacing a subgoal with one that is - logically stronger (rather than equivalent). For example, a - subgoal involving equality on a quotient type could be replaced - with a subgoal involving equality (instead of the corresponding - equivalence relation) on the underlying raw type. - - \item @{method (HOL) "transfer_prover"} method assists with proving - a transfer rule for a new constant, provided the constant is - defined in terms of other constants that already have transfer - rules. It should be applied after unfolding the constant - definitions. - - \item @{attribute (HOL) "transfer_rule"} attribute maintains a - collection of transfer rules, which relate constants at two - different types. Typical transfer rules may relate different type - instances of the same polymorphic constant, or they may relate an - operation on a raw type to a corresponding operation on an - abstract type (quotient or subtype). For example: - - @{text "((A ===> B) ===> list_all2 A ===> list_all2 B) map map"}\\ - @{text "(cr_int ===> cr_int ===> cr_int) (\(x,y) (u,v). (x+u, y+v)) plus"} - - Lemmas involving predicates on relations can also be registered - using the same attribute. For example: - - @{text "bi_unique A \ (list_all2 A ===> op =) distinct distinct"}\\ - @{text "\bi_unique A; bi_unique B\ \ bi_unique (prod_rel A B)"} - - \item @{attribute (HOL) relator_eq} attribute collects identity laws - for relators of various type constructors, e.g. @{text "list_all2 - (op =) = (op =)"}. The @{method (HOL) transfer} method uses these - lemmas to infer transfer rules for non-polymorphic constants on - the fly. - - \end{description} - -*} - - -section {* Lifting package *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "setup_lifting"} & : & @{text "local_theory \ local_theory"}\\ - @{command_def (HOL) "lift_definition"} & : & @{text "local_theory \ proof(prove)"}\\ - @{command_def (HOL) "print_quotmaps"} & : & @{text "context \"}\\ - @{command_def (HOL) "print_quotients"} & : & @{text "context \"}\\ - @{attribute_def (HOL) "quot_map"} & : & @{text attribute} \\ - @{attribute_def (HOL) "invariant_commute"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{command (HOL) setup_lifting} ('(' 'no_abs_code' ')')? \\ - @{syntax thmref} @{syntax thmref}?; - "} - - @{rail " - @@{command (HOL) lift_definition} @{syntax name} '::' @{syntax type} @{syntax mixfix}? \\ - 'is' @{syntax term}; - "} - - \begin{description} - - \item @{command (HOL) "setup_lifting"} Sets up the Lifting package - to work with a user-defined type. The user must provide either a - quotient theorem @{text "Quotient R Abs Rep T"} or a - type_definition theorem @{text "type_definition Rep Abs A"}. The - package configures transfer rules for equality and quantifiers on - the type, and sets up the @{command_def (HOL) "lift_definition"} - command to work with the type. In the case of a quotient theorem, - an optional theorem @{text "reflp R"} can be provided as a second - argument. This allows the package to generate stronger transfer - rules. - - @{command (HOL) "setup_lifting"} is called automatically if a - quotient type is defined by the command @{command (HOL) - "quotient_type"} from the Quotient package. - - If @{command (HOL) "setup_lifting"} is called with a - type_definition theorem, the abstract type implicitly defined by - the theorem is declared as an abstract type in the code - generator. This allows @{command (HOL) "lift_definition"} to - register (generated) code certificate theorems as abstract code - equations in the code generator. The option @{text "no_abs_code"} - of the command @{command (HOL) "setup_lifting"} can turn off that - behavior and causes that code certificate theorems generated by - @{command (HOL) "lift_definition"} are not registred as abstract - code equations. - - \item @{command (HOL) "lift_definition"} @{text "f :: \ is t"} - Defines a new function @{text f} with an abstract type @{text \} - in terms of a corresponding operation @{text t} on a - representation type. The term @{text t} doesn't have to be - necessarily a constant but it can be any term. - - Users must discharge a respectfulness proof obligation when each - constant is defined. For a type copy, i.e. a typedef with @{text - UNIV}, the proof is discharged automatically. The obligation is - presented in a user-friendly, readable form. A respectfulness - theorem in the standard format @{text f.rsp} and a transfer rule - @{text f.tranfer} for the Transfer package are generated by the - package. - - Integration with code_abstype: For typedefs (e.g. subtypes - corresponding to a datatype invariant, such as dlist), @{command - (HOL) "lift_definition"} generates a code certificate theorem - @{text f.rep_eq} and sets up code generation for each constant. - - \item @{command (HOL) "print_quotmaps"} prints stored quotient map - theorems. - - \item @{command (HOL) "print_quotients"} prints stored quotient - theorems. - - \item @{attribute (HOL) quot_map} registers a quotient map - theorem. For examples see @{file - "~~/src/HOL/Library/Quotient_List.thy"} or other Quotient_*.thy - files. - - \item @{attribute (HOL) invariant_commute} registers a theorem which - shows a relationship between the constant @{text - Lifting.invariant} (used for internal encoding of proper subtypes) - and a relator. Such theorems allows the package to hide @{text - Lifting.invariant} from a user in a user-readable form of a - respectfulness theorem. For examples see @{file - "~~/src/HOL/Library/Quotient_List.thy"} or other Quotient_*.thy - files. - - \end{description} -*} - - -section {* Quotient types *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "quotient_type"} & : & @{text "local_theory \ proof(prove)"}\\ - @{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \ proof(prove)"}\\ - @{command_def (HOL) "print_quotmapsQ3"} & : & @{text "context \"}\\ - @{command_def (HOL) "print_quotientsQ3"} & : & @{text "context \"}\\ - @{command_def (HOL) "print_quotconsts"} & : & @{text "context \"}\\ - @{method_def (HOL) "lifting"} & : & @{text method} \\ - @{method_def (HOL) "lifting_setup"} & : & @{text method} \\ - @{method_def (HOL) "descending"} & : & @{text method} \\ - @{method_def (HOL) "descending_setup"} & : & @{text method} \\ - @{method_def (HOL) "partiality_descending"} & : & @{text method} \\ - @{method_def (HOL) "partiality_descending_setup"} & : & @{text method} \\ - @{method_def (HOL) "regularize"} & : & @{text method} \\ - @{method_def (HOL) "injection"} & : & @{text method} \\ - @{method_def (HOL) "cleaning"} & : & @{text method} \\ - @{attribute_def (HOL) "quot_thm"} & : & @{text attribute} \\ - @{attribute_def (HOL) "quot_lifted"} & : & @{text attribute} \\ - @{attribute_def (HOL) "quot_respect"} & : & @{text attribute} \\ - @{attribute_def (HOL) "quot_preserve"} & : & @{text attribute} \\ - \end{matharray} - - The quotient package defines a new quotient type given a raw type - and a partial equivalence relation. It also includes automation for - transporting definitions and theorems. It can automatically produce - definitions and theorems on the quotient type, given the - corresponding constants and facts on the raw type. - - @{rail " - @@{command (HOL) quotient_type} (spec + @'and'); - - spec: @{syntax typespec} @{syntax mixfix}? '=' \\ - @{syntax type} '/' ('partial' ':')? @{syntax term} \\ - (@'morphisms' @{syntax name} @{syntax name})?; - "} - - @{rail " - @@{command (HOL) quotient_definition} constdecl? @{syntax thmdecl}? \\ - @{syntax term} 'is' @{syntax term}; - - constdecl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? - "} - - @{rail " - @@{method (HOL) lifting} @{syntax thmrefs}? - ; - - @@{method (HOL) lifting_setup} @{syntax thmrefs}? - ; - "} - - \begin{description} - - \item @{command (HOL) "quotient_type"} defines quotient types. The - injection from a quotient type to a raw type is called @{text - rep_t}, its inverse @{text abs_t} unless explicit @{keyword (HOL) - "morphisms"} specification provides alternative names. @{command - (HOL) "quotient_type"} requires the user to prove that the relation - is an equivalence relation (predicate @{text equivp}), unless the - user specifies explicitely @{text partial} in which case the - obligation is @{text part_equivp}. A quotient defined with @{text - partial} is weaker in the sense that less things can be proved - automatically. - - \item @{command (HOL) "quotient_definition"} defines a constant on - the quotient type. - - \item @{command (HOL) "print_quotmapsQ3"} prints quotient map - functions. - - \item @{command (HOL) "print_quotientsQ3"} prints quotients. - - \item @{command (HOL) "print_quotconsts"} prints quotient constants. - - \item @{method (HOL) "lifting"} and @{method (HOL) "lifting_setup"} - methods match the current goal with the given raw theorem to be - lifted producing three new subgoals: regularization, injection and - cleaning subgoals. @{method (HOL) "lifting"} tries to apply the - heuristics for automatically solving these three subgoals and - leaves only the subgoals unsolved by the heuristics to the user as - opposed to @{method (HOL) "lifting_setup"} which leaves the three - subgoals unsolved. - - \item @{method (HOL) "descending"} and @{method (HOL) - "descending_setup"} try to guess a raw statement that would lift - to the current subgoal. Such statement is assumed as a new subgoal - and @{method (HOL) "descending"} continues in the same way as - @{method (HOL) "lifting"} does. @{method (HOL) "descending"} tries - to solve the arising regularization, injection and cleaning - subgoals with the analoguous method @{method (HOL) - "descending_setup"} which leaves the four unsolved subgoals. - - \item @{method (HOL) "partiality_descending"} finds the regularized - theorem that would lift to the current subgoal, lifts it and - leaves as a subgoal. This method can be used with partial - equivalence quotients where the non regularized statements would - not be true. @{method (HOL) "partiality_descending_setup"} leaves - the injection and cleaning subgoals unchanged. - - \item @{method (HOL) "regularize"} applies the regularization - heuristics to the current subgoal. - - \item @{method (HOL) "injection"} applies the injection heuristics - to the current goal using the stored quotient respectfulness - theorems. - - \item @{method (HOL) "cleaning"} applies the injection cleaning - heuristics to the current subgoal using the stored quotient - preservation theorems. - - \item @{attribute (HOL) quot_lifted} attribute tries to - automatically transport the theorem to the quotient type. - The attribute uses all the defined quotients types and quotient - constants often producing undesired results or theorems that - cannot be lifted. - - \item @{attribute (HOL) quot_respect} and @{attribute (HOL) - quot_preserve} attributes declare a theorem as a respectfulness - and preservation theorem respectively. These are stored in the - local theory store and used by the @{method (HOL) "injection"} - and @{method (HOL) "cleaning"} methods respectively. - - \item @{attribute (HOL) quot_thm} declares that a certain theorem - is a quotient extension theorem. Quotient extension theorems - allow for quotienting inside container types. Given a polymorphic - type that serves as a container, a map function defined for this - container using @{command (HOL) "enriched_type"} and a relation - map defined for for the container type, the quotient extension - theorem should be @{term "Quotient3 R Abs Rep \ Quotient3 - (rel_map R) (map Abs) (map Rep)"}. Quotient extension theorems - are stored in a database and are used all the steps of lifting - theorems. - - \end{description} -*} - - -section {* Coercive subtyping *} - -text {* - \begin{matharray}{rcl} - @{attribute_def (HOL) coercion} & : & @{text attribute} \\ - @{attribute_def (HOL) coercion_enabled} & : & @{text attribute} \\ - @{attribute_def (HOL) coercion_map} & : & @{text attribute} \\ - \end{matharray} - - Coercive subtyping allows the user to omit explicit type - conversions, also called \emph{coercions}. Type inference will add - them as necessary when parsing a term. See - \cite{traytel-berghofer-nipkow-2011} for details. - - @{rail " - @@{attribute (HOL) coercion} (@{syntax term})? - ; - "} - @{rail " - @@{attribute (HOL) coercion_map} (@{syntax term})? - ; - "} - - \begin{description} - - \item @{attribute (HOL) "coercion"}~@{text "f"} registers a new - coercion function @{text "f :: \\<^isub>1 \ \\<^isub>2"} where @{text "\\<^isub>1"} and - @{text "\\<^isub>2"} are type constructors without arguments. Coercions are - composed by the inference algorithm if needed. Note that the type - inference algorithm is complete only if the registered coercions - form a lattice. - - \item @{attribute (HOL) "coercion_map"}~@{text "map"} registers a - new map function to lift coercions through type constructors. The - function @{text "map"} must conform to the following type pattern - - \begin{matharray}{lll} - @{text "map"} & @{text "::"} & - @{text "f\<^isub>1 \ \ \ f\<^isub>n \ (\\<^isub>1, \, \\<^isub>n) t \ (\\<^isub>1, \, \\<^isub>n) t"} \\ - \end{matharray} - - where @{text "t"} is a type constructor and @{text "f\<^isub>i"} is of type - @{text "\\<^isub>i \ \\<^isub>i"} or @{text "\\<^isub>i \ \\<^isub>i"}. Registering a map function - overwrites any existing map function for this particular type - constructor. - - \item @{attribute (HOL) "coercion_enabled"} enables the coercion - inference algorithm. - - \end{description} -*} - - -section {* Arithmetic proof support *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) arith} & : & @{text method} \\ - @{attribute_def (HOL) arith} & : & @{text attribute} \\ - @{attribute_def (HOL) arith_split} & : & @{text attribute} \\ - \end{matharray} - - \begin{description} - - \item @{method (HOL) arith} decides linear arithmetic problems (on - types @{text nat}, @{text int}, @{text real}). Any current facts - are inserted into the goal before running the procedure. - - \item @{attribute (HOL) arith} declares facts that are supplied to - the arithmetic provers implicitly. - - \item @{attribute (HOL) arith_split} attribute declares case split - rules to be expanded before @{method (HOL) arith} is invoked. - - \end{description} - - Note that a simpler (but faster) arithmetic prover is already - invoked by the Simplifier. -*} - - -section {* Intuitionistic proof search *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) iprover} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method (HOL) iprover} ( @{syntax rulemod} * ) - "} - - \begin{description} - - \item @{method (HOL) iprover} performs intuitionistic proof search, - depending on specifically declared rules from the context, or given - as explicit arguments. Chained facts are inserted into the goal - before commencing proof search. - - Rules need to be classified as @{attribute (Pure) intro}, - @{attribute (Pure) elim}, or @{attribute (Pure) dest}; here the - ``@{text "!"}'' indicator refers to ``safe'' rules, which may be - applied aggressively (without considering back-tracking later). - Rules declared with ``@{text "?"}'' are ignored in proof search (the - single-step @{method (Pure) rule} method still observes these). An - explicit weight annotation may be given as well; otherwise the - number of rule premises will be taken into account here. - - \end{description} -*} - - -section {* Model Elimination and Resolution *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) "meson"} & : & @{text method} \\ - @{method_def (HOL) "metis"} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method (HOL) meson} @{syntax thmrefs}? - ; - - @@{method (HOL) metis} - ('(' ('partial_types' | 'full_types' | 'no_types' | @{syntax name}) ')')? - @{syntax thmrefs}? - "} - - \begin{description} - - \item @{method (HOL) meson} implements Loveland's model elimination - procedure \cite{loveland-78}. See @{file - "~~/src/HOL/ex/Meson_Test.thy"} for examples. - - \item @{method (HOL) metis} combines ordered resolution and ordered - paramodulation to find first-order (or mildly higher-order) proofs. - The first optional argument specifies a type encoding; see the - Sledgehammer manual \cite{isabelle-sledgehammer} for details. The - directory @{file "~~/src/HOL/Metis_Examples"} contains several small - theories developed to a large extent using @{method (HOL) metis}. - - \end{description} -*} - - -section {* Coherent Logic *} - -text {* - \begin{matharray}{rcl} - @{method_def (HOL) "coherent"} & : & @{text method} \\ - \end{matharray} - - @{rail " - @@{method (HOL) coherent} @{syntax thmrefs}? - "} - - \begin{description} - - \item @{method (HOL) coherent} solves problems of \emph{Coherent - Logic} \cite{Bezem-Coquand:2005}, which covers applications in - confluence theory, lattice theory and projective geometry. See - @{file "~~/src/HOL/ex/Coherent.thy"} for some examples. - - \end{description} -*} - - -section {* Proving propositions *} - -text {* - In addition to the standard proof methods, a number of diagnosis - tools search for proofs and provide an Isar proof snippet on success. - These tools are available via the following commands. - - \begin{matharray}{rcl} - @{command_def (HOL) "solve_direct"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "try"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "try0"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "sledgehammer"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "sledgehammer_params"} & : & @{text "theory \ theory"} - \end{matharray} - - @{rail " - @@{command (HOL) try} - ; - - @@{command (HOL) try0} ( ( ( 'simp' | 'intro' | 'elim' | 'dest' ) ':' @{syntax thmrefs} ) + ) ? - @{syntax nat}? - ; - - @@{command (HOL) sledgehammer} ( '[' args ']' )? facts? @{syntax nat}? - ; - - @@{command (HOL) sledgehammer_params} ( ( '[' args ']' ) ? ) - ; - - args: ( @{syntax name} '=' value + ',' ) - ; - - facts: '(' ( ( ( ( 'add' | 'del' ) ':' ) ? @{syntax thmrefs} ) + ) ? ')' - ; - "} % FIXME check args "value" - - \begin{description} - - \item @{command (HOL) "solve_direct"} checks whether the current - subgoals can be solved directly by an existing theorem. Duplicate - lemmas can be detected in this way. - - \item @{command (HOL) "try0"} attempts to prove a subgoal - using a combination of standard proof methods (@{method auto}, - @{method simp}, @{method blast}, etc.). Additional facts supplied - via @{text "simp:"}, @{text "intro:"}, @{text "elim:"}, and @{text - "dest:"} are passed to the appropriate proof methods. - - \item @{command (HOL) "try"} attempts to prove or disprove a subgoal - using a combination of provers and disprovers (@{command (HOL) - "solve_direct"}, @{command (HOL) "quickcheck"}, @{command (HOL) - "try0"}, @{command (HOL) "sledgehammer"}, @{command (HOL) - "nitpick"}). - - \item @{command (HOL) "sledgehammer"} attempts to prove a subgoal - using external automatic provers (resolution provers and SMT - solvers). See the Sledgehammer manual \cite{isabelle-sledgehammer} - for details. - - \item @{command (HOL) "sledgehammer_params"} changes @{command (HOL) - "sledgehammer"} configuration options persistently. - - \end{description} -*} - - -section {* Checking and refuting propositions *} - -text {* - Identifying incorrect propositions usually involves evaluation of - particular assignments and systematic counterexample search. This - is supported by the following commands. - - \begin{matharray}{rcl} - @{command_def (HOL) "value"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def (HOL) "values"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def (HOL) "quickcheck"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "refute"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "nitpick"}@{text "\<^sup>*"} & : & @{text "proof \"} \\ - @{command_def (HOL) "quickcheck_params"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "refute_params"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "nitpick_params"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "quickcheck_generator"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "find_unused_assms"} & : & @{text "context \"} - \end{matharray} - - @{rail " - @@{command (HOL) value} ( '[' @{syntax name} ']' )? modes? @{syntax term} - ; - - @@{command (HOL) values} modes? @{syntax nat}? @{syntax term} - ; - - (@@{command (HOL) quickcheck} | @@{command (HOL) refute} | @@{command (HOL) nitpick}) - ( '[' args ']' )? @{syntax nat}? - ; - - (@@{command (HOL) quickcheck_params} | @@{command (HOL) refute_params} | - @@{command (HOL) nitpick_params}) ( '[' args ']' )? - ; - - @@{command (HOL) quickcheck_generator} @{syntax nameref} \\ - 'operations:' ( @{syntax term} +) - ; - - @@{command (HOL) find_unused_assms} @{syntax name}? - ; - - modes: '(' (@{syntax name} +) ')' - ; - - args: ( @{syntax name} '=' value + ',' ) - ; - "} % FIXME check "value" - - \begin{description} - - \item @{command (HOL) "value"}~@{text t} evaluates and prints a - term; optionally @{text modes} can be specified, which are appended - to the current print mode; see \secref{sec:print-modes}. - Internally, the evaluation is performed by registered evaluators, - which are invoked sequentially until a result is returned. - Alternatively a specific evaluator can be selected using square - brackets; typical evaluators use the current set of code equations - to normalize and include @{text simp} for fully symbolic evaluation - using the simplifier, @{text nbe} for \emph{normalization by - evaluation} and \emph{code} for code generation in SML. - - \item @{command (HOL) "values"}~@{text t} enumerates a set - comprehension by evaluation and prints its values up to the given - number of solutions; optionally @{text modes} can be specified, - which are appended to the current print mode; see - \secref{sec:print-modes}. - - \item @{command (HOL) "quickcheck"} tests the current goal for - counterexamples using a series of assignments for its free - variables; by default the first subgoal is tested, an other can be - selected explicitly using an optional goal index. Assignments can - be chosen exhausting the search space upto a given size, or using a - fixed number of random assignments in the search space, or exploring - the search space symbolically using narrowing. By default, - quickcheck uses exhaustive testing. A number of configuration - options are supported for @{command (HOL) "quickcheck"}, notably: - - \begin{description} - - \item[@{text tester}] specifies which testing approach to apply. - There are three testers, @{text exhaustive}, @{text random}, and - @{text narrowing}. An unknown configuration option is treated as - an argument to tester, making @{text "tester ="} optional. When - multiple testers are given, these are applied in parallel. If no - tester is specified, quickcheck uses the testers that are set - active, i.e., configurations @{attribute - quickcheck_exhaustive_active}, @{attribute - quickcheck_random_active}, @{attribute - quickcheck_narrowing_active} are set to true. - - \item[@{text size}] specifies the maximum size of the search space - for assignment values. - - \item[@{text genuine_only}] sets quickcheck only to return genuine - counterexample, but not potentially spurious counterexamples due - to underspecified functions. - - \item[@{text abort_potential}] sets quickcheck to abort once it - found a potentially spurious counterexample and to not continue - to search for a further genuine counterexample. - For this option to be effective, the @{text genuine_only} option - must be set to false. - - \item[@{text eval}] takes a term or a list of terms and evaluates - these terms under the variable assignment found by quickcheck. - This option is currently only supported by the default - (exhaustive) tester. - - \item[@{text iterations}] sets how many sets of assignments are - generated for each particular size. - - \item[@{text no_assms}] specifies whether assumptions in - structured proofs should be ignored. - - \item[@{text locale}] specifies how to process conjectures in - a locale context, i.e., they can be interpreted or expanded. - The option is a whitespace-separated list of the two words - @{text interpret} and @{text expand}. The list determines the - order they are employed. The default setting is to first use - interpretations and then test the expanded conjecture. - The option is only provided as attribute declaration, but not - as parameter to the command. - - \item[@{text timeout}] sets the time limit in seconds. - - \item[@{text default_type}] sets the type(s) generally used to - instantiate type variables. - - \item[@{text report}] if set quickcheck reports how many tests - fulfilled the preconditions. - - \item[@{text use_subtype}] if set quickcheck automatically lifts - conjectures to registered subtypes if possible, and tests the - lifted conjecture. - - \item[@{text quiet}] if set quickcheck does not output anything - while testing. - - \item[@{text verbose}] if set quickcheck informs about the current - size and cardinality while testing. - - \item[@{text expect}] can be used to check if the user's - expectation was met (@{text no_expectation}, @{text - no_counterexample}, or @{text counterexample}). - - \end{description} - - These option can be given within square brackets. - - \item @{command (HOL) "quickcheck_params"} changes @{command (HOL) - "quickcheck"} configuration options persistently. - - \item @{command (HOL) "quickcheck_generator"} creates random and - exhaustive value generators for a given type and operations. It - generates values by using the operations as if they were - constructors of that type. - - \item @{command (HOL) "refute"} tests the current goal for - counterexamples using a reduction to SAT. The following - configuration options are supported: - - \begin{description} - - \item[@{text minsize}] specifies the minimum size (cardinality) of - the models to search for. - - \item[@{text maxsize}] specifies the maximum size (cardinality) of - the models to search for. Nonpositive values mean @{text "\"}. - - \item[@{text maxvars}] specifies the maximum number of Boolean - variables to use when transforming the term into a propositional - formula. Nonpositive values mean @{text "\"}. - - \item[@{text satsolver}] specifies the SAT solver to use. - - \item[@{text no_assms}] specifies whether assumptions in - structured proofs should be ignored. - - \item[@{text maxtime}] sets the time limit in seconds. - - \item[@{text expect}] can be used to check if the user's - expectation was met (@{text genuine}, @{text potential}, @{text - none}, or @{text unknown}). - - \end{description} - - These option can be given within square brackets. - - \item @{command (HOL) "refute_params"} changes @{command (HOL) - "refute"} configuration options persistently. - - \item @{command (HOL) "nitpick"} tests the current goal for - counterexamples using a reduction to first-order relational - logic. See the Nitpick manual \cite{isabelle-nitpick} for details. - - \item @{command (HOL) "nitpick_params"} changes @{command (HOL) - "nitpick"} configuration options persistently. - - \item @{command (HOL) "find_unused_assms"} finds potentially superfluous - assumptions in theorems using quickcheck. - It takes the theory name to be checked for superfluous assumptions as - optional argument. If not provided, it checks the current theory. - Options to the internal quickcheck invocations can be changed with - common configuration declarations. - - \end{description} -*} - - -section {* Unstructured case analysis and induction \label{sec:hol-induct-tac} *} - -text {* - The following tools of Isabelle/HOL support cases analysis and - induction in unstructured tactic scripts; see also - \secref{sec:cases-induct} for proper Isar versions of similar ideas. - - \begin{matharray}{rcl} - @{method_def (HOL) case_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def (HOL) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\ - @{method_def (HOL) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\ - @{command_def (HOL) "inductive_cases"}@{text "\<^sup>*"} & : & @{text "local_theory \ local_theory"} \\ - \end{matharray} - - @{rail " - @@{method (HOL) case_tac} @{syntax goal_spec}? @{syntax term} rule? - ; - @@{method (HOL) induct_tac} @{syntax goal_spec}? (@{syntax insts} * @'and') rule? - ; - @@{method (HOL) ind_cases} (@{syntax prop}+) (@'for' (@{syntax name}+))? - ; - @@{command (HOL) inductive_cases} (@{syntax thmdecl}? (@{syntax prop}+) + @'and') - ; - - rule: 'rule' ':' @{syntax thmref} - "} - - \begin{description} - - \item @{method (HOL) case_tac} and @{method (HOL) induct_tac} admit - to reason about inductive types. Rules are selected according to - the declarations by the @{attribute cases} and @{attribute induct} - attributes, cf.\ \secref{sec:cases-induct}. The @{command (HOL) - datatype} package already takes care of this. - - These unstructured tactics feature both goal addressing and dynamic - instantiation. Note that named rule cases are \emph{not} provided - as would be by the proper @{method cases} and @{method induct} proof - methods (see \secref{sec:cases-induct}). Unlike the @{method - induct} method, @{method induct_tac} does not handle structured rule - statements, only the compact object-logic conclusion of the subgoal - being addressed. - - \item @{method (HOL) ind_cases} and @{command (HOL) - "inductive_cases"} provide an interface to the internal @{ML_text - mk_cases} operation. Rules are simplified in an unrestricted - forward manner. - - While @{method (HOL) ind_cases} is a proof method to apply the - result immediately as elimination rules, @{command (HOL) - "inductive_cases"} provides case split theorems at the theory level - for later use. The @{keyword "for"} argument of the @{method (HOL) - ind_cases} method allows to specify a list of variables that should - be generalized before applying the resulting rule. - - \end{description} -*} - - -section {* Executable code *} - -text {* For validation purposes, it is often useful to \emph{execute} - specifications. In principle, execution could be simulated by - Isabelle's inference kernel, i.e. by a combination of resolution and - simplification. Unfortunately, this approach is rather inefficient. - A more efficient way of executing specifications is to translate - them into a functional programming language such as ML. - - Isabelle provides a generic framework to support code generation - from executable specifications. Isabelle/HOL instantiates these - mechanisms in a way that is amenable to end-user applications. Code - can be generated for functional programs (including overloading - using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml}, - Haskell \cite{haskell-revised-report} and Scala - \cite{scala-overview-tech-report}. Conceptually, code generation is - split up in three steps: \emph{selection} of code theorems, - \emph{translation} into an abstract executable view and - \emph{serialization} to a specific \emph{target language}. - Inductive specifications can be executed using the predicate - compiler which operates within HOL. See \cite{isabelle-codegen} for - an introduction. - - \begin{matharray}{rcl} - @{command_def (HOL) "export_code"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def (HOL) code} & : & @{text attribute} \\ - @{command_def (HOL) "code_abort"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_datatype"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "print_codesetup"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def (HOL) code_unfold} & : & @{text attribute} \\ - @{attribute_def (HOL) code_post} & : & @{text attribute} \\ - @{attribute_def (HOL) code_abbrev} & : & @{text attribute} \\ - @{command_def (HOL) "print_codeproc"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def (HOL) "code_thms"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def (HOL) "code_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def (HOL) "code_const"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_type"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_class"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_instance"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_reserved"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_monad"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_include"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_modulename"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_reflect"} & : & @{text "theory \ theory"} \\ - @{command_def (HOL) "code_pred"} & : & @{text "theory \ proof(prove)"} - \end{matharray} - - @{rail " - @@{command (HOL) export_code} ( constexpr + ) \\ - ( ( @'in' target ( @'module_name' @{syntax string} ) ? \\ - ( @'file' ( @{syntax string} | '-' ) ) ? ( '(' args ')' ) ?) + ) ? - ; - - const: @{syntax term} - ; - - constexpr: ( const | 'name._' | '_' ) - ; - - typeconstructor: @{syntax nameref} - ; - - class: @{syntax nameref} - ; - - target: 'SML' | 'OCaml' | 'Haskell' | 'Scala' - ; - - @@{attribute (HOL) code} ( 'del' | 'abstype' | 'abstract' )? - ; - - @@{command (HOL) code_abort} ( const + ) - ; - - @@{command (HOL) code_datatype} ( const + ) - ; - - @@{attribute (HOL) code_unfold} ( 'del' ) ? - ; - - @@{attribute (HOL) code_post} ( 'del' ) ? - ; - - @@{attribute (HOL) code_abbrev} - ; - - @@{command (HOL) code_thms} ( constexpr + ) ? - ; - - @@{command (HOL) code_deps} ( constexpr + ) ? - ; - - @@{command (HOL) code_const} (const + @'and') \\ - ( ( '(' target ( syntax ? + @'and' ) ')' ) + ) - ; - - @@{command (HOL) code_type} (typeconstructor + @'and') \\ - ( ( '(' target ( syntax ? + @'and' ) ')' ) + ) - ; - - @@{command (HOL) code_class} (class + @'and') \\ - ( ( '(' target \\ ( @{syntax string} ? + @'and' ) ')' ) + ) - ; - - @@{command (HOL) code_instance} (( typeconstructor '::' class ) + @'and') \\ - ( ( '(' target ( '-' ? + @'and' ) ')' ) + ) - ; - - @@{command (HOL) code_reserved} target ( @{syntax string} + ) - ; - - @@{command (HOL) code_monad} const const target - ; - - @@{command (HOL) code_include} target ( @{syntax string} ( @{syntax string} | '-') ) - ; - - @@{command (HOL) code_modulename} target ( ( @{syntax string} @{syntax string} ) + ) - ; - - @@{command (HOL) code_reflect} @{syntax string} \\ - ( @'datatypes' ( @{syntax string} '=' ( '_' | ( @{syntax string} + '|' ) + @'and' ) ) ) ? \\ - ( @'functions' ( @{syntax string} + ) ) ? ( @'file' @{syntax string} ) ? - ; - - @@{command (HOL) code_pred} \\( '(' @'modes' ':' modedecl ')')? \\ const - ; - - syntax: @{syntax string} | ( @'infix' | @'infixl' | @'infixr' ) @{syntax nat} @{syntax string} - ; - - modedecl: (modes | ((const ':' modes) \\ - (@'and' ((const ':' modes @'and') +))?)) - ; - - modes: mode @'as' const - "} - - \begin{description} - - \item @{command (HOL) "export_code"} generates code for a given list - of constants in the specified target language(s). If no - serialization instruction is given, only abstract code is generated - internally. - - Constants may be specified by giving them literally, referring to - all executable contants within a certain theory by giving @{text - "name.*"}, or referring to \emph{all} executable constants currently - available by giving @{text "*"}. - - By default, for each involved theory one corresponding name space - module is generated. Alternativly, a module name may be specified - after the @{keyword "module_name"} keyword; then \emph{all} code is - placed in this module. - - For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification - refers to a single file; for \emph{Haskell}, it refers to a whole - directory, where code is generated in multiple files reflecting the - module hierarchy. Omitting the file specification denotes standard - output. - - Serializers take an optional list of arguments in parentheses. For - \emph{SML} and \emph{OCaml}, ``@{text no_signatures}`` omits - explicit module signatures. - - For \emph{Haskell} a module name prefix may be given using the - ``@{text "root:"}'' argument; ``@{text string_classes}'' adds a - ``@{verbatim "deriving (Read, Show)"}'' clause to each appropriate - datatype declaration. - - \item @{attribute (HOL) code} explicitly selects (or with option - ``@{text "del"}'' deselects) a code equation for code generation. - Usually packages introducing code equations provide a reasonable - default setup for selection. Variants @{text "code abstype"} and - @{text "code abstract"} declare abstract datatype certificates or - code equations on abstract datatype representations respectively. - - \item @{command (HOL) "code_abort"} declares constants which are not - required to have a definition by means of code equations; if needed - these are implemented by program abort instead. - - \item @{command (HOL) "code_datatype"} specifies a constructor set - for a logical type. - - \item @{command (HOL) "print_codesetup"} gives an overview on - selected code equations and code generator datatypes. - - \item @{attribute (HOL) code_unfold} declares (or with option - ``@{text "del"}'' removes) theorems which during preprocessing - are applied as rewrite rules to any code equation or evaluation - input. - - \item @{attribute (HOL) code_post} declares (or with option ``@{text - "del"}'' removes) theorems which are applied as rewrite rules to any - result of an evaluation. - - \item @{attribute (HOL) code_abbrev} declares equations which are - applied as rewrite rules to any result of an evaluation and - symmetrically during preprocessing to any code equation or evaluation - input. - - \item @{command (HOL) "print_codeproc"} prints the setup of the code - generator preprocessor. - - \item @{command (HOL) "code_thms"} prints a list of theorems - representing the corresponding program containing all given - constants after preprocessing. - - \item @{command (HOL) "code_deps"} visualizes dependencies of - theorems representing the corresponding program containing all given - constants after preprocessing. - - \item @{command (HOL) "code_const"} associates a list of constants - with target-specific serializations; omitting a serialization - deletes an existing serialization. - - \item @{command (HOL) "code_type"} associates a list of type - constructors with target-specific serializations; omitting a - serialization deletes an existing serialization. - - \item @{command (HOL) "code_class"} associates a list of classes - with target-specific class names; omitting a serialization deletes - an existing serialization. This applies only to \emph{Haskell}. - - \item @{command (HOL) "code_instance"} declares a list of type - constructor / class instance relations as ``already present'' for a - given target. Omitting a ``@{text "-"}'' deletes an existing - ``already present'' declaration. This applies only to - \emph{Haskell}. - - \item @{command (HOL) "code_reserved"} declares a list of names as - reserved for a given target, preventing it to be shadowed by any - generated code. - - \item @{command (HOL) "code_monad"} provides an auxiliary mechanism - to generate monadic code for Haskell. - - \item @{command (HOL) "code_include"} adds arbitrary named content - (``include'') to generated code. A ``@{text "-"}'' as last argument - will remove an already added ``include''. - - \item @{command (HOL) "code_modulename"} declares aliasings from one - module name onto another. - - \item @{command (HOL) "code_reflect"} without a ``@{text "file"}'' - argument compiles code into the system runtime environment and - modifies the code generator setup that future invocations of system - runtime code generation referring to one of the ``@{text - "datatypes"}'' or ``@{text "functions"}'' entities use these - precompiled entities. With a ``@{text "file"}'' argument, the - corresponding code is generated into that specified file without - modifying the code generator setup. - - \item @{command (HOL) "code_pred"} creates code equations for a - predicate given a set of introduction rules. Optional mode - annotations determine which arguments are supposed to be input or - output. If alternative introduction rules are declared, one must - prove a corresponding elimination rule. - - \end{description} -*} - - -section {* Definition by specification \label{sec:hol-specification} *} - -text {* - \begin{matharray}{rcl} - @{command_def (HOL) "specification"} & : & @{text "theory \ proof(prove)"} \\ - @{command_def (HOL) "ax_specification"} & : & @{text "theory \ proof(prove)"} \\ - \end{matharray} - - @{rail " - (@@{command (HOL) specification} | @@{command (HOL) ax_specification}) - '(' (decl +) ')' \\ (@{syntax thmdecl}? @{syntax prop} +) - ; - decl: ((@{syntax name} ':')? @{syntax term} '(' @'overloaded' ')'?) - "} - - \begin{description} - - \item @{command (HOL) "specification"}~@{text "decls \"} sets up a - goal stating the existence of terms with the properties specified to - hold for the constants given in @{text decls}. After finishing the - proof, the theory will be augmented with definitions for the given - constants, as well as with theorems stating the properties for these - constants. - - \item @{command (HOL) "ax_specification"}~@{text "decls \"} sets up - a goal stating the existence of terms with the properties specified - to hold for the constants given in @{text decls}. After finishing - the proof, the theory will be augmented with axioms expressing the - properties given in the first place. - - \item @{text decl} declares a constant to be defined by the - specification given. The definition for the constant @{text c} is - bound to the name @{text c_def} unless a theorem name is given in - the declaration. Overloaded constants should be declared as such. - - \end{description} - - Whether to use @{command (HOL) "specification"} or @{command (HOL) - "ax_specification"} is to some extent a matter of style. @{command - (HOL) "specification"} introduces no new axioms, and so by - construction cannot introduce inconsistencies, whereas @{command - (HOL) "ax_specification"} does introduce axioms, but only after the - user has explicitly proven it to be safe. A practical issue must be - considered, though: After introducing two constants with the same - properties using @{command (HOL) "specification"}, one can prove - that the two constants are, in fact, equal. If this might be a - problem, one should use @{command (HOL) "ax_specification"}. -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Inner_Syntax.thy --- a/doc-src/IsarRef/Thy/Inner_Syntax.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1574 +0,0 @@ -theory Inner_Syntax -imports Base Main -begin - -chapter {* Inner syntax --- the term language \label{ch:inner-syntax} *} - -text {* The inner syntax of Isabelle provides concrete notation for - the main entities of the logical framework, notably @{text - "\"}-terms with types and type classes. Applications may either - extend existing syntactic categories by additional notation, or - define new sub-languages that are linked to the standard term - language via some explicit markers. For example @{verbatim - FOO}~@{text "foo"} could embed the syntax corresponding for some - user-defined nonterminal @{text "foo"} --- within the bounds of the - given lexical syntax of Isabelle/Pure. - - The most basic way to specify concrete syntax for logical entities - works via mixfix annotations (\secref{sec:mixfix}), which may be - usually given as part of the original declaration or via explicit - notation commands later on (\secref{sec:notation}). This already - covers many needs of concrete syntax without having to understand - the full complexity of inner syntax layers. - - Further details of the syntax engine involves the classical - distinction of lexical language versus context-free grammar (see - \secref{sec:pure-syntax}), and various mechanisms for \emph{syntax - transformations} (see \secref{sec:syntax-transformations}). -*} - - -section {* Printing logical entities *} - -subsection {* Diagnostic commands \label{sec:print-diag} *} - -text {* - \begin{matharray}{rcl} - @{command_def "typ"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "term"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "prop"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "thm"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "prf"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "full_prf"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "pr"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - \end{matharray} - - These diagnostic commands assist interactive development by printing - internal logical entities in a human-readable fashion. - - @{rail " - @@{command typ} @{syntax modes}? @{syntax type} ('::' @{syntax sort})? - ; - @@{command term} @{syntax modes}? @{syntax term} - ; - @@{command prop} @{syntax modes}? @{syntax prop} - ; - @@{command thm} @{syntax modes}? @{syntax thmrefs} - ; - ( @@{command prf} | @@{command full_prf} ) @{syntax modes}? @{syntax thmrefs}? - ; - @@{command pr} @{syntax modes}? @{syntax nat}? - ; - - @{syntax_def modes}: '(' (@{syntax name} + ) ')' - "} - - \begin{description} - - \item @{command "typ"}~@{text \} reads and prints a type expression - according to the current context. - - \item @{command "typ"}~@{text "\ :: s"} uses type-inference to - determine the most general way to make @{text "\"} conform to sort - @{text "s"}. For concrete @{text "\"} this checks if the type - belongs to that sort. Dummy type parameters ``@{text "_"}'' - (underscore) are assigned to fresh type variables with most general - sorts, according the the principles of type-inference. - - \item @{command "term"}~@{text t} and @{command "prop"}~@{text \} - read, type-check and print terms or propositions according to the - current theory or proof context; the inferred type of @{text t} is - output as well. Note that these commands are also useful in - inspecting the current environment of term abbreviations. - - \item @{command "thm"}~@{text "a\<^sub>1 \ a\<^sub>n"} retrieves - theorems from the current theory or proof context. Note that any - attributes included in the theorem specifications are applied to a - temporary context derived from the current theory or proof; the - result is discarded, i.e.\ attributes involved in @{text "a\<^sub>1, - \, a\<^sub>n"} do not have any permanent effect. - - \item @{command "prf"} displays the (compact) proof term of the - current proof state (if present), or of the given theorems. Note - that this requires proof terms to be switched on for the current - object logic (see the ``Proof terms'' section of the Isabelle - reference manual for information on how to do this). - - \item @{command "full_prf"} is like @{command "prf"}, but displays - the full proof term, i.e.\ also displays information omitted in the - compact proof term, which is denoted by ``@{text _}'' placeholders - there. - - \item @{command "pr"}~@{text "goals"} prints the current proof state - (if present), including current facts and goals. The optional limit - arguments affect the number of goals to be displayed, which is - initially 10. Omitting limit value leaves the current setting - unchanged. - - \end{description} - - All of the diagnostic commands above admit a list of @{text modes} - to be specified, which is appended to the current print mode; see - also \secref{sec:print-modes}. Thus the output behavior may be - modified according particular print mode features. For example, - @{command "pr"}~@{text "(latex xsymbols)"} would print the current - proof state with mathematical symbols and special characters - represented in {\LaTeX} source, according to the Isabelle style - \cite{isabelle-sys}. - - Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more - systematic way to include formal items into the printed text - document. -*} - - -subsection {* Details of printed content *} - -text {* - \begin{tabular}{rcll} - @{attribute_def show_types} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_sorts} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_consts} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_abbrevs} & : & @{text attribute} & default @{text true} \\ - @{attribute_def show_brackets} & : & @{text attribute} & default @{text false} \\ - @{attribute_def names_long} & : & @{text attribute} & default @{text false} \\ - @{attribute_def names_short} & : & @{text attribute} & default @{text false} \\ - @{attribute_def names_unique} & : & @{text attribute} & default @{text true} \\ - @{attribute_def eta_contract} & : & @{text attribute} & default @{text true} \\ - @{attribute_def goals_limit} & : & @{text attribute} & default @{text 10} \\ - @{attribute_def show_main_goal} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_hyps} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_tags} & : & @{text attribute} & default @{text false} \\ - @{attribute_def show_question_marks} & : & @{text attribute} & default @{text true} \\ - \end{tabular} - \medskip - - These configuration options control the detail of information that - is displayed for types, terms, theorems, goals etc. See also - \secref{sec:config}. - - \begin{description} - - \item @{attribute show_types} and @{attribute show_sorts} control - printing of type constraints for term variables, and sort - constraints for type variables. By default, neither of these are - shown in output. If @{attribute show_sorts} is enabled, types are - always shown as well. - - Note that displaying types and sorts may explain why a polymorphic - inference rule fails to resolve with some goal, or why a rewrite - rule does not apply as expected. - - \item @{attribute show_consts} controls printing of types of - constants when displaying a goal state. - - Note that the output can be enormous, because polymorphic constants - often occur at several different type instances. - - \item @{attribute show_abbrevs} controls folding of constant - abbreviations. - - \item @{attribute show_brackets} controls bracketing in pretty - printed output. If enabled, all sub-expressions of the pretty - printing tree will be parenthesized, even if this produces malformed - term syntax! This crude way of showing the internal structure of - pretty printed entities may occasionally help to diagnose problems - with operator priorities, for example. - - \item @{attribute names_long}, @{attribute names_short}, and - @{attribute names_unique} control the way of printing fully - qualified internal names in external form. See also - \secref{sec:antiq} for the document antiquotation options of the - same names. - - \item @{attribute eta_contract} controls @{text "\"}-contracted - printing of terms. - - The @{text \}-contraction law asserts @{prop "(\x. f x) \ f"}, - provided @{text x} is not free in @{text f}. It asserts - \emph{extensionality} of functions: @{prop "f \ g"} if @{prop "f x \ - g x"} for all @{text x}. Higher-order unification frequently puts - terms into a fully @{text \}-expanded form. For example, if @{text - F} has type @{text "(\ \ \) \ \"} then its expanded form is @{term - "\h. F (\x. h x)"}. - - Enabling @{attribute eta_contract} makes Isabelle perform @{text - \}-contractions before printing, so that @{term "\h. F (\x. h x)"} - appears simply as @{text F}. - - Note that the distinction between a term and its @{text \}-expanded - form occasionally matters. While higher-order resolution and - rewriting operate modulo @{text "\\\"}-conversion, some other tools - might look at terms more discretely. - - \item @{attribute goals_limit} controls the maximum number of - subgoals to be shown in goal output. - - \item @{attribute show_main_goal} controls whether the main result - to be proven should be displayed. This information might be - relevant for schematic goals, to inspect the current claim that has - been synthesized so far. - - \item @{attribute show_hyps} controls printing of implicit - hypotheses of local facts. Normally, only those hypotheses are - displayed that are \emph{not} covered by the assumptions of the - current context: this situation indicates a fault in some tool being - used. - - By enabling @{attribute show_hyps}, output of \emph{all} hypotheses - can be enforced, which is occasionally useful for diagnostic - purposes. - - \item @{attribute show_tags} controls printing of extra annotations - within theorems, such as internal position information, or the case - names being attached by the attribute @{attribute case_names}. - - Note that the @{attribute tagged} and @{attribute untagged} - attributes provide low-level access to the collection of tags - associated with a theorem. - - \item @{attribute show_question_marks} controls printing of question - marks for schematic variables, such as @{text ?x}. Only the leading - question mark is affected, the remaining text is unchanged - (including proper markup for schematic variables that might be - relevant for user interfaces). - - \end{description} -*} - - -subsection {* Alternative print modes \label{sec:print-modes} *} - -text {* - \begin{mldecls} - @{index_ML print_mode_value: "unit -> string list"} \\ - @{index_ML Print_Mode.with_modes: "string list -> ('a -> 'b) -> 'a -> 'b"} \\ - \end{mldecls} - - The \emph{print mode} facility allows to modify various operations - for printing. Commands like @{command typ}, @{command term}, - @{command thm} (see \secref{sec:print-diag}) take additional print - modes as optional argument. The underlying ML operations are as - follows. - - \begin{description} - - \item @{ML "print_mode_value ()"} yields the list of currently - active print mode names. This should be understood as symbolic - representation of certain individual features for printing (with - precedence from left to right). - - \item @{ML Print_Mode.with_modes}~@{text "modes f x"} evaluates - @{text "f x"} in an execution context where the print mode is - prepended by the given @{text "modes"}. This provides a thread-safe - way to augment print modes. It is also monotonic in the set of mode - names: it retains the default print mode that certain - user-interfaces might have installed for their proper functioning! - - \end{description} - - \begin{warn} - The old global reference @{ML print_mode} should never be used - directly in applications. Its main reason for being publicly - accessible is to support historic versions of Proof~General. - \end{warn} - - \medskip The pretty printer for inner syntax maintains alternative - mixfix productions for any print mode name invented by the user, say - in commands like @{command notation} or @{command abbreviation}. - Mode names can be arbitrary, but the following ones have a specific - meaning by convention: - - \begin{itemize} - - \item @{verbatim "\"\""} (the empty string): default mode; - implicitly active as last element in the list of modes. - - \item @{verbatim input}: dummy print mode that is never active; may - be used to specify notation that is only available for input. - - \item @{verbatim internal} dummy print mode that is never active; - used internally in Isabelle/Pure. - - \item @{verbatim xsymbols}: enable proper mathematical symbols - instead of ASCII art.\footnote{This traditional mode name stems from - the ``X-Symbol'' package for old versions Proof~General with XEmacs, - although that package has been superseded by Unicode in recent - years.} - - \item @{verbatim HTML}: additional mode that is active in HTML - presentation of Isabelle theory sources; allows to provide - alternative output notation. - - \item @{verbatim latex}: additional mode that is active in {\LaTeX} - document preparation of Isabelle theory sources; allows to provide - alternative output notation. - - \end{itemize} -*} - - -subsection {* Printing limits *} - -text {* - \begin{mldecls} - @{index_ML Pretty.margin_default: "int Unsynchronized.ref"} \\ - @{index_ML print_depth: "int -> unit"} \\ - \end{mldecls} - - These ML functions set limits for pretty printed text. - - \begin{description} - - \item @{ML Pretty.margin_default} indicates the global default for - the right margin of the built-in pretty printer, with initial value - 76. Note that user-interfaces typically control margins - automatically when resizing windows, or even bypass the formatting - engine of Isabelle/ML altogether and do it within the front end via - Isabelle/Scala. - - \item @{ML print_depth}~@{text n} limits the printing depth of the - ML toplevel pretty printer; the precise effect depends on the ML - compiler and run-time system. Typically @{text n} should be less - than 10. Bigger values such as 100--1000 are useful for debugging. - - \end{description} -*} - - -section {* Mixfix annotations \label{sec:mixfix} *} - -text {* Mixfix annotations specify concrete \emph{inner syntax} of - Isabelle types and terms. Locally fixed parameters in toplevel - theorem statements, locale and class specifications also admit - mixfix annotations in a fairly uniform manner. A mixfix annotation - describes describes the concrete syntax, the translation to abstract - syntax, and the pretty printing. Special case annotations provide a - simple means of specifying infix operators and binders. - - Isabelle mixfix syntax is inspired by {\OBJ} \cite{OBJ}. It allows - to specify any context-free priority grammar, which is more general - than the fixity declarations of ML and Prolog. - - @{rail " - @{syntax_def mixfix}: '(' mfix ')' - ; - @{syntax_def struct_mixfix}: '(' ( mfix | @'structure' ) ')' - ; - - mfix: @{syntax template} prios? @{syntax nat}? | - (@'infix' | @'infixl' | @'infixr') @{syntax template} @{syntax nat} | - @'binder' @{syntax template} prios? @{syntax nat} - ; - template: string - ; - prios: '[' (@{syntax nat} + ',') ']' - "} - - The string given as @{text template} may include literal text, - spacing, blocks, and arguments (denoted by ``@{text _}''); the - special symbol ``@{verbatim "\"}'' (printed as ``@{text "\"}'') - represents an index argument that specifies an implicit structure - reference (see also \secref{sec:locale}). Infix and binder - declarations provide common abbreviations for particular mixfix - declarations. So in practice, mixfix templates mostly degenerate to - literal text for concrete syntax, such as ``@{verbatim "++"}'' for - an infix symbol. -*} - - -subsection {* The general mixfix form *} - -text {* In full generality, mixfix declarations work as follows. - Suppose a constant @{text "c :: \\<^sub>1 \ \ \\<^sub>n \ \"} is annotated by - @{text "(mixfix [p\<^sub>1, \, p\<^sub>n] p)"}, where @{text "mixfix"} is a string - @{text "d\<^sub>0 _ d\<^sub>1 _ \ _ d\<^sub>n"} consisting of delimiters that surround - argument positions as indicated by underscores. - - Altogether this determines a production for a context-free priority - grammar, where for each argument @{text "i"} the syntactic category - is determined by @{text "\\<^sub>i"} (with priority @{text "p\<^sub>i"}), and the - result category is determined from @{text "\"} (with priority @{text - "p"}). Priority specifications are optional, with default 0 for - arguments and 1000 for the result.\footnote{Omitting priorities is - prone to syntactic ambiguities unless the delimiter tokens determine - fully bracketed notation, as in @{text "if _ then _ else _ fi"}.} - - Since @{text "\"} may be again a function type, the constant - type scheme may have more argument positions than the mixfix - pattern. Printing a nested application @{text "c t\<^sub>1 \ t\<^sub>m"} for - @{text "m > n"} works by attaching concrete notation only to the - innermost part, essentially by printing @{text "(c t\<^sub>1 \ t\<^sub>n) \ t\<^sub>m"} - instead. If a term has fewer arguments than specified in the mixfix - template, the concrete syntax is ignored. - - \medskip A mixfix template may also contain additional directives - for pretty printing, notably spaces, blocks, and breaks. The - general template format is a sequence over any of the following - entities. - - \begin{description} - - \item @{text "d"} is a delimiter, namely a non-empty sequence of - characters other than the following special characters: - - \smallskip - \begin{tabular}{ll} - @{verbatim "'"} & single quote \\ - @{verbatim "_"} & underscore \\ - @{text "\"} & index symbol \\ - @{verbatim "("} & open parenthesis \\ - @{verbatim ")"} & close parenthesis \\ - @{verbatim "/"} & slash \\ - \end{tabular} - \medskip - - \item @{verbatim "'"} escapes the special meaning of these - meta-characters, producing a literal version of the following - character, unless that is a blank. - - A single quote followed by a blank separates delimiters, without - affecting printing, but input tokens may have additional white space - here. - - \item @{verbatim "_"} is an argument position, which stands for a - certain syntactic category in the underlying grammar. - - \item @{text "\"} is an indexed argument position; this is the place - where implicit structure arguments can be attached. - - \item @{text "s"} is a non-empty sequence of spaces for printing. - This and the following specifications do not affect parsing at all. - - \item @{verbatim "("}@{text n} opens a pretty printing block. The - optional number specifies how much indentation to add when a line - break occurs within the block. If the parenthesis is not followed - by digits, the indentation defaults to 0. A block specified via - @{verbatim "(00"} is unbreakable. - - \item @{verbatim ")"} closes a pretty printing block. - - \item @{verbatim "//"} forces a line break. - - \item @{verbatim "/"}@{text s} allows a line break. Here @{text s} - stands for the string of spaces (zero or more) right after the - slash. These spaces are printed if the break is \emph{not} taken. - - \end{description} - - The general idea of pretty printing with blocks and breaks is also - described in \cite{paulson-ml2}; it goes back to \cite{Oppen:1980}. -*} - - -subsection {* Infixes *} - -text {* Infix operators are specified by convenient short forms that - abbreviate general mixfix annotations as follows: - - \begin{center} - \begin{tabular}{lll} - - @{verbatim "("}@{keyword_def "infix"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} - & @{text "\"} & - @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p + 1"}@{verbatim ", "}@{text "p + 1"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ - @{verbatim "("}@{keyword_def "infixl"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} - & @{text "\"} & - @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p"}@{verbatim ", "}@{text "p + 1"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ - @{verbatim "("}@{keyword_def "infixr"}~@{verbatim "\""}@{text sy}@{verbatim "\""} @{text "p"}@{verbatim ")"} - & @{text "\"} & - @{verbatim "(\"(_ "}@{text sy}@{verbatim "/ _)\" ["}@{text "p + 1"}@{verbatim ", "}@{text "p"}@{verbatim "]"}@{text " p"}@{verbatim ")"} \\ - - \end{tabular} - \end{center} - - The mixfix template @{verbatim "\"(_ "}@{text sy}@{verbatim "/ _)\""} - specifies two argument positions; the delimiter is preceded by a - space and followed by a space or line break; the entire phrase is a - pretty printing block. - - The alternative notation @{verbatim "op"}~@{text sy} is introduced - in addition. Thus any infix operator may be written in prefix form - (as in ML), independently of the number of arguments in the term. -*} - - -subsection {* Binders *} - -text {* A \emph{binder} is a variable-binding construct such as a - quantifier. The idea to formalize @{text "\x. b"} as @{text "All - (\x. b)"} for @{text "All :: ('a \ bool) \ bool"} already goes back - to \cite{church40}. Isabelle declarations of certain higher-order - operators may be annotated with @{keyword_def "binder"} annotations - as follows: - - \begin{center} - @{text "c :: "}@{verbatim "\""}@{text "(\\<^sub>1 \ \\<^sub>2) \ \\<^sub>3"}@{verbatim "\" ("}@{keyword "binder"}@{verbatim " \""}@{text "sy"}@{verbatim "\" ["}@{text "p"}@{verbatim "] "}@{text "q"}@{verbatim ")"} - \end{center} - - This introduces concrete binder syntax @{text "sy x. b"}, where - @{text x} is a bound variable of type @{text "\\<^sub>1"}, the body @{text - b} has type @{text "\\<^sub>2"} and the whole term has type @{text "\\<^sub>3"}. - The optional integer @{text p} specifies the syntactic priority of - the body; the default is @{text "q"}, which is also the priority of - the whole construct. - - Internally, the binder syntax is expanded to something like this: - \begin{center} - @{text "c_binder :: "}@{verbatim "\""}@{text "idts \ \\<^sub>2 \ \\<^sub>3"}@{verbatim "\" (\"(3"}@{text sy}@{verbatim "_./ _)\" [0, "}@{text "p"}@{verbatim "] "}@{text "q"}@{verbatim ")"} - \end{center} - - Here @{syntax (inner) idts} is the nonterminal symbol for a list of - identifiers with optional type constraints (see also - \secref{sec:pure-grammar}). The mixfix template @{verbatim - "\"(3"}@{text sy}@{verbatim "_./ _)\""} defines argument positions - for the bound identifiers and the body, separated by a dot with - optional line break; the entire phrase is a pretty printing block of - indentation level 3. Note that there is no extra space after @{text - "sy"}, so it needs to be included user specification if the binder - syntax ends with a token that may be continued by an identifier - token at the start of @{syntax (inner) idts}. - - Furthermore, a syntax translation to transforms @{text "c_binder x\<^sub>1 - \ x\<^sub>n b"} into iterated application @{text "c (\x\<^sub>1. \ c (\x\<^sub>n. b)\)"}. - This works in both directions, for parsing and printing. *} - - -section {* Explicit notation \label{sec:notation} *} - -text {* - \begin{matharray}{rcll} - @{command_def "type_notation"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "no_type_notation"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "notation"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "no_notation"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "write"} & : & @{text "proof(state) \ proof(state)"} \\ - \end{matharray} - - Commands that introduce new logical entities (terms or types) - usually allow to provide mixfix annotations on the spot, which is - convenient for default notation. Nonetheless, the syntax may be - modified later on by declarations for explicit notation. This - allows to add or delete mixfix annotations for of existing logical - entities within the current context. - - @{rail " - (@@{command type_notation} | @@{command no_type_notation}) @{syntax target}? - @{syntax mode}? \\ (@{syntax nameref} @{syntax mixfix} + @'and') - ; - (@@{command notation} | @@{command no_notation}) @{syntax target}? @{syntax mode}? \\ - (@{syntax nameref} @{syntax struct_mixfix} + @'and') - ; - @@{command write} @{syntax mode}? (@{syntax nameref} @{syntax struct_mixfix} + @'and') - "} - - \begin{description} - - \item @{command "type_notation"}~@{text "c (mx)"} associates mixfix - syntax with an existing type constructor. The arity of the - constructor is retrieved from the context. - - \item @{command "no_type_notation"} is similar to @{command - "type_notation"}, but removes the specified syntax annotation from - the present context. - - \item @{command "notation"}~@{text "c (mx)"} associates mixfix - syntax with an existing constant or fixed variable. The type - declaration of the given entity is retrieved from the context. - - \item @{command "no_notation"} is similar to @{command "notation"}, - but removes the specified syntax annotation from the present - context. - - \item @{command "write"} is similar to @{command "notation"}, but - works within an Isar proof body. - - \end{description} -*} - - -section {* The Pure syntax \label{sec:pure-syntax} *} - -subsection {* Lexical matters \label{sec:inner-lex} *} - -text {* The inner lexical syntax vaguely resembles the outer one - (\secref{sec:outer-lex}), but some details are different. There are - two main categories of inner syntax tokens: - - \begin{enumerate} - - \item \emph{delimiters} --- the literal tokens occurring in - productions of the given priority grammar (cf.\ - \secref{sec:priority-grammar}); - - \item \emph{named tokens} --- various categories of identifiers etc. - - \end{enumerate} - - Delimiters override named tokens and may thus render certain - identifiers inaccessible. Sometimes the logical context admits - alternative ways to refer to the same entity, potentially via - qualified names. - - \medskip The categories for named tokens are defined once and for - all as follows, reusing some categories of the outer token syntax - (\secref{sec:outer-lex}). - - \begin{center} - \begin{supertabular}{rcl} - @{syntax_def (inner) id} & = & @{syntax_ref ident} \\ - @{syntax_def (inner) longid} & = & @{syntax_ref longident} \\ - @{syntax_def (inner) var} & = & @{syntax_ref var} \\ - @{syntax_def (inner) tid} & = & @{syntax_ref typefree} \\ - @{syntax_def (inner) tvar} & = & @{syntax_ref typevar} \\ - @{syntax_def (inner) num_token} & = & @{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat} \\ - @{syntax_def (inner) float_token} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ - @{syntax_def (inner) xnum_token} & = & @{verbatim "#"}@{syntax_ref nat}@{text " | "}@{verbatim "#-"}@{syntax_ref nat} \\ - - @{syntax_def (inner) str_token} & = & @{verbatim "''"} @{text "\"} @{verbatim "''"} \\ - \end{supertabular} - \end{center} - - The token categories @{syntax (inner) num_token}, @{syntax (inner) - float_token}, @{syntax (inner) xnum_token}, and @{syntax (inner) - str_token} are not used in Pure. Object-logics may implement numerals - and string constants by adding appropriate syntax declarations, - together with some translation functions (e.g.\ see Isabelle/HOL). - - The derived categories @{syntax_def (inner) num_const}, @{syntax_def - (inner) float_const}, and @{syntax_def (inner) num_const} provide - robust access to the respective tokens: the syntax tree holds a - syntactic constant instead of a free variable. -*} - - -subsection {* Priority grammars \label{sec:priority-grammar} *} - -text {* A context-free grammar consists of a set of \emph{terminal - symbols}, a set of \emph{nonterminal symbols} and a set of - \emph{productions}. Productions have the form @{text "A = \"}, - where @{text A} is a nonterminal and @{text \} is a string of - terminals and nonterminals. One designated nonterminal is called - the \emph{root symbol}. The language defined by the grammar - consists of all strings of terminals that can be derived from the - root symbol by applying productions as rewrite rules. - - The standard Isabelle parser for inner syntax uses a \emph{priority - grammar}. Each nonterminal is decorated by an integer priority: - @{text "A\<^sup>(\<^sup>p\<^sup>)"}. In a derivation, @{text "A\<^sup>(\<^sup>p\<^sup>)"} may be rewritten - using a production @{text "A\<^sup>(\<^sup>q\<^sup>) = \"} only if @{text "p \ q"}. Any - priority grammar can be translated into a normal context-free - grammar by introducing new nonterminals and productions. - - \medskip Formally, a set of context free productions @{text G} - induces a derivation relation @{text "\\<^sub>G"} as follows. Let @{text - \} and @{text \} denote strings of terminal or nonterminal symbols. - Then @{text "\ A\<^sup>(\<^sup>p\<^sup>) \ \\<^sub>G \ \ \"} holds if and only if @{text G} - contains some production @{text "A\<^sup>(\<^sup>q\<^sup>) = \"} for @{text "p \ q"}. - - \medskip The following grammar for arithmetic expressions - demonstrates how binding power and associativity of operators can be - enforced by priorities. - - \begin{center} - \begin{tabular}{rclr} - @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim "("} @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim ")"} \\ - @{text "A\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "="} & @{verbatim 0} \\ - @{text "A\<^sup>(\<^sup>0\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} \\ - @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "="} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} \\ - @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "="} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} \\ - \end{tabular} - \end{center} - The choice of priorities determines that @{verbatim "-"} binds - tighter than @{verbatim "*"}, which binds tighter than @{verbatim - "+"}. Furthermore @{verbatim "+"} associates to the left and - @{verbatim "*"} to the right. - - \medskip For clarity, grammars obey these conventions: - \begin{itemize} - - \item All priorities must lie between 0 and 1000. - - \item Priority 0 on the right-hand side and priority 1000 on the - left-hand side may be omitted. - - \item The production @{text "A\<^sup>(\<^sup>p\<^sup>) = \"} is written as @{text "A = \ - (p)"}, i.e.\ the priority of the left-hand side actually appears in - a column on the far right. - - \item Alternatives are separated by @{text "|"}. - - \item Repetition is indicated by dots @{text "(\)"} in an informal - but obvious way. - - \end{itemize} - - Using these conventions, the example grammar specification above - takes the form: - \begin{center} - \begin{tabular}{rclc} - @{text A} & @{text "="} & @{verbatim "("} @{text A} @{verbatim ")"} \\ - & @{text "|"} & @{verbatim 0} & \qquad\qquad \\ - & @{text "|"} & @{text A} @{verbatim "+"} @{text "A\<^sup>(\<^sup>1\<^sup>)"} & @{text "(0)"} \\ - & @{text "|"} & @{text "A\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "*"} @{text "A\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ - & @{text "|"} & @{verbatim "-"} @{text "A\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ - \end{tabular} - \end{center} -*} - - -subsection {* The Pure grammar \label{sec:pure-grammar} *} - -text {* The priority grammar of the @{text "Pure"} theory is defined - approximately like this: - - \begin{center} - \begin{supertabular}{rclr} - - @{syntax_def (inner) any} & = & @{text "prop | logic"} \\\\ - - @{syntax_def (inner) prop} & = & @{verbatim "("} @{text prop} @{verbatim ")"} \\ - & @{text "|"} & @{text "prop\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ - & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "=="} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ - & @{text "|"} & @{text "any\<^sup>(\<^sup>3\<^sup>)"} @{text "\"} @{text "any\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ - & @{text "|"} & @{text "prop\<^sup>(\<^sup>3\<^sup>)"} @{verbatim "&&&"} @{text "prop\<^sup>(\<^sup>2\<^sup>)"} & @{text "(2)"} \\ - & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ - & @{text "|"} & @{text "prop\<^sup>(\<^sup>2\<^sup>)"} @{text "\"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ - & @{text "|"} & @{verbatim "[|"} @{text prop} @{verbatim ";"} @{text "\"} @{verbatim ";"} @{text prop} @{verbatim "|]"} @{verbatim "==>"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ - & @{text "|"} & @{text "\"} @{text prop} @{verbatim ";"} @{text "\"} @{verbatim ";"} @{text prop} @{text "\"} @{text "\"} @{text "prop\<^sup>(\<^sup>1\<^sup>)"} & @{text "(1)"} \\ - & @{text "|"} & @{verbatim "!!"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ - & @{text "|"} & @{text "\"} @{text idts} @{verbatim "."} @{text prop} & @{text "(0)"} \\ - & @{text "|"} & @{verbatim OFCLASS} @{verbatim "("} @{text type} @{verbatim ","} @{text logic} @{verbatim ")"} \\ - & @{text "|"} & @{verbatim SORT_CONSTRAINT} @{verbatim "("} @{text type} @{verbatim ")"} \\ - & @{text "|"} & @{verbatim TERM} @{text logic} \\ - & @{text "|"} & @{verbatim PROP} @{text aprop} \\\\ - - @{syntax_def (inner) aprop} & = & @{verbatim "("} @{text aprop} @{verbatim ")"} \\ - & @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ - & @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ - & @{text "|"} & @{verbatim XCONST} @{text "id | "}@{verbatim XCONST} @{text "longid"} \\ - & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\\\ - - @{syntax_def (inner) logic} & = & @{verbatim "("} @{text logic} @{verbatim ")"} \\ - & @{text "|"} & @{text "logic\<^sup>(\<^sup>4\<^sup>)"} @{verbatim "::"} @{text type} & @{text "(3)"} \\ - & @{text "|"} & @{text "id | longid | var | "}@{verbatim "_"}@{text " | "}@{verbatim "..."} \\ - & @{text "|"} & @{verbatim CONST} @{text "id | "}@{verbatim CONST} @{text "longid"} \\ - & @{text "|"} & @{verbatim XCONST} @{text "id | "}@{verbatim XCONST} @{text "longid"} \\ - & @{text "|"} & @{text "logic\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) \ any\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} & @{text "(999)"} \\ - & @{text "|"} & @{text "\ index\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>)"} \\ - & @{text "|"} & @{verbatim "%"} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ - & @{text "|"} & @{text \} @{text pttrns} @{verbatim "."} @{text "any\<^sup>(\<^sup>3\<^sup>)"} & @{text "(3)"} \\ - & @{text "|"} & @{verbatim op} @{verbatim "=="}@{text " | "}@{verbatim op} @{text "\"}@{text " | "}@{verbatim op} @{verbatim "&&&"} \\ - & @{text "|"} & @{verbatim op} @{verbatim "==>"}@{text " | "}@{verbatim op} @{text "\"} \\ - & @{text "|"} & @{verbatim TYPE} @{verbatim "("} @{text type} @{verbatim ")"} \\\\ - - @{syntax_def (inner) idt} & = & @{verbatim "("} @{text idt} @{verbatim ")"}@{text " | id | "}@{verbatim "_"} \\ - & @{text "|"} & @{text id} @{verbatim "::"} @{text type} & @{text "(0)"} \\ - & @{text "|"} & @{verbatim "_"} @{verbatim "::"} @{text type} & @{text "(0)"} \\\\ - - @{syntax_def (inner) index} & = & @{verbatim "\<^bsub>"} @{text "logic\<^sup>(\<^sup>0\<^sup>)"} @{verbatim "\<^esub>"}@{text " | | \"} \\\\ - - @{syntax_def (inner) idts} & = & @{text "idt | idt\<^sup>(\<^sup>1\<^sup>) idts"} & @{text "(0)"} \\\\ - - @{syntax_def (inner) pttrn} & = & @{text idt} \\\\ - - @{syntax_def (inner) pttrns} & = & @{text "pttrn | pttrn\<^sup>(\<^sup>1\<^sup>) pttrns"} & @{text "(0)"} \\\\ - - @{syntax_def (inner) type} & = & @{verbatim "("} @{text type} @{verbatim ")"} \\ - & @{text "|"} & @{text "tid | tvar | "}@{verbatim "_"} \\ - & @{text "|"} & @{text "tid"} @{verbatim "::"} @{text "sort | tvar "}@{verbatim "::"} @{text "sort | "}@{verbatim "_"} @{verbatim "::"} @{text "sort"} \\ - & @{text "|"} & @{text "type_name | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) type_name"} \\ - & @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim ")"} @{text type_name} \\ - & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ - & @{text "|"} & @{text "type\<^sup>(\<^sup>1\<^sup>)"} @{text "\"} @{text type} & @{text "(0)"} \\ - & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim "]"} @{verbatim "=>"} @{text type} & @{text "(0)"} \\ - & @{text "|"} & @{verbatim "["} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim "]"} @{text "\"} @{text type} & @{text "(0)"} \\ - @{syntax_def (inner) type_name} & = & @{text "id | longid"} \\\\ - - @{syntax_def (inner) sort} & = & @{syntax class_name}~@{text " | "}@{verbatim "{}"} \\ - & @{text "|"} & @{verbatim "{"} @{syntax class_name} @{verbatim ","} @{text "\"} @{verbatim ","} @{syntax class_name} @{verbatim "}"} \\ - @{syntax_def (inner) class_name} & = & @{text "id | longid"} \\ - \end{supertabular} - \end{center} - - \medskip Here literal terminals are printed @{verbatim "verbatim"}; - see also \secref{sec:inner-lex} for further token categories of the - inner syntax. The meaning of the nonterminals defined by the above - grammar is as follows: - - \begin{description} - - \item @{syntax_ref (inner) any} denotes any term. - - \item @{syntax_ref (inner) prop} denotes meta-level propositions, - which are terms of type @{typ prop}. The syntax of such formulae of - the meta-logic is carefully distinguished from usual conventions for - object-logics. In particular, plain @{text "\"}-term notation is - \emph{not} recognized as @{syntax (inner) prop}. - - \item @{syntax_ref (inner) aprop} denotes atomic propositions, which - are embedded into regular @{syntax (inner) prop} by means of an - explicit @{verbatim PROP} token. - - Terms of type @{typ prop} with non-constant head, e.g.\ a plain - variable, are printed in this form. Constants that yield type @{typ - prop} are expected to provide their own concrete syntax; otherwise - the printed version will appear like @{syntax (inner) logic} and - cannot be parsed again as @{syntax (inner) prop}. - - \item @{syntax_ref (inner) logic} denotes arbitrary terms of a - logical type, excluding type @{typ prop}. This is the main - syntactic category of object-logic entities, covering plain @{text - \}-term notation (variables, abstraction, application), plus - anything defined by the user. - - When specifying notation for logical entities, all logical types - (excluding @{typ prop}) are \emph{collapsed} to this single category - of @{syntax (inner) logic}. - - \item @{syntax_ref (inner) index} denotes an optional index term for - indexed syntax. If omitted, it refers to the first @{keyword - "structure"} variable in the context. The special dummy ``@{text - "\"}'' serves as pattern variable in mixfix annotations that - introduce indexed notation. - - \item @{syntax_ref (inner) idt} denotes identifiers, possibly - constrained by types. - - \item @{syntax_ref (inner) idts} denotes a sequence of @{syntax_ref - (inner) idt}. This is the most basic category for variables in - iterated binders, such as @{text "\"} or @{text "\"}. - - \item @{syntax_ref (inner) pttrn} and @{syntax_ref (inner) pttrns} - denote patterns for abstraction, cases bindings etc. In Pure, these - categories start as a merely copy of @{syntax (inner) idt} and - @{syntax (inner) idts}, respectively. Object-logics may add - additional productions for binding forms. - - \item @{syntax_ref (inner) type} denotes types of the meta-logic. - - \item @{syntax_ref (inner) sort} denotes meta-level sorts. - - \end{description} - - Here are some further explanations of certain syntax features. - - \begin{itemize} - - \item In @{syntax (inner) idts}, note that @{text "x :: nat y"} is - parsed as @{text "x :: (nat y)"}, treating @{text y} like a type - constructor applied to @{text nat}. To avoid this interpretation, - write @{text "(x :: nat) y"} with explicit parentheses. - - \item Similarly, @{text "x :: nat y :: nat"} is parsed as @{text "x :: - (nat y :: nat)"}. The correct form is @{text "(x :: nat) (y :: - nat)"}, or @{text "(x :: nat) y :: nat"} if @{text y} is last in the - sequence of identifiers. - - \item Type constraints for terms bind very weakly. For example, - @{text "x < y :: nat"} is normally parsed as @{text "(x < y) :: - nat"}, unless @{text "<"} has a very low priority, in which case the - input is likely to be ambiguous. The correct form is @{text "x < (y - :: nat)"}. - - \item Constraints may be either written with two literal colons - ``@{verbatim "::"}'' or the double-colon symbol @{verbatim "\"}, - which actually looks exactly the same in some {\LaTeX} styles. - - \item Dummy variables (written as underscore) may occur in different - roles. - - \begin{description} - - \item A type ``@{text "_"}'' or ``@{text "_ :: sort"}'' acts like an - anonymous inference parameter, which is filled-in according to the - most general type produced by the type-checking phase. - - \item A bound ``@{text "_"}'' refers to a vacuous abstraction, where - the body does not refer to the binding introduced here. As in the - term @{term "\x _. x"}, which is @{text "\"}-equivalent to @{text - "\x y. x"}. - - \item A free ``@{text "_"}'' refers to an implicit outer binding. - Higher definitional packages usually allow forms like @{text "f x _ - = x"}. - - \item A schematic ``@{text "_"}'' (within a term pattern, see - \secref{sec:term-decls}) refers to an anonymous variable that is - implicitly abstracted over its context of locally bound variables. - For example, this allows pattern matching of @{text "{x. f x = g - x}"} against @{text "{x. _ = _}"}, or even @{text "{_. _ = _}"} by - using both bound and schematic dummies. - - \end{description} - - \item The three literal dots ``@{verbatim "..."}'' may be also - written as ellipsis symbol @{verbatim "\"}. In both cases this - refers to a special schematic variable, which is bound in the - context. This special term abbreviation works nicely with - calculational reasoning (\secref{sec:calculation}). - - \item @{verbatim CONST} ensures that the given identifier is treated - as constant term, and passed through the parse tree in fully - internalized form. This is particularly relevant for translation - rules (\secref{sec:syn-trans}), notably on the RHS. - - \item @{verbatim XCONST} is similar to @{verbatim CONST}, but - retains the constant name as given. This is only relevant to - translation rules (\secref{sec:syn-trans}), notably on the LHS. - - \end{itemize} -*} - - -subsection {* Inspecting the syntax *} - -text {* - \begin{matharray}{rcl} - @{command_def "print_syntax"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - \end{matharray} - - \begin{description} - - \item @{command "print_syntax"} prints the inner syntax of the - current context. The output can be quite large; the most important - sections are explained below. - - \begin{description} - - \item @{text "lexicon"} lists the delimiters of the inner token - language; see \secref{sec:inner-lex}. - - \item @{text "prods"} lists the productions of the underlying - priority grammar; see \secref{sec:priority-grammar}. - - The nonterminal @{text "A\<^sup>(\<^sup>p\<^sup>)"} is rendered in plain text as @{text - "A[p]"}; delimiters are quoted. Many productions have an extra - @{text "\ => name"}. These names later become the heads of parse - trees; they also guide the pretty printer. - - Productions without such parse tree names are called \emph{copy - productions}. Their right-hand side must have exactly one - nonterminal symbol (or named token). The parser does not create a - new parse tree node for copy productions, but simply returns the - parse tree of the right-hand symbol. - - If the right-hand side of a copy production consists of a single - nonterminal without any delimiters, then it is called a \emph{chain - production}. Chain productions act as abbreviations: conceptually, - they are removed from the grammar by adding new productions. - Priority information attached to chain productions is ignored; only - the dummy value @{text "-1"} is displayed. - - \item @{text "print modes"} lists the alternative print modes - provided by this grammar; see \secref{sec:print-modes}. - - \item @{text "parse_rules"} and @{text "print_rules"} relate to - syntax translations (macros); see \secref{sec:syn-trans}. - - \item @{text "parse_ast_translation"} and @{text - "print_ast_translation"} list sets of constants that invoke - translation functions for abstract syntax trees, which are only - required in very special situations; see \secref{sec:tr-funs}. - - \item @{text "parse_translation"} and @{text "print_translation"} - list the sets of constants that invoke regular translation - functions; see \secref{sec:tr-funs}. - - \end{description} - - \end{description} -*} - - -subsection {* Ambiguity of parsed expressions *} - -text {* - \begin{tabular}{rcll} - @{attribute_def syntax_ambiguity_warning} & : & @{text attribute} & default @{text true} \\ - @{attribute_def syntax_ambiguity_limit} & : & @{text attribute} & default @{text 10} \\ - \end{tabular} - - Depending on the grammar and the given input, parsing may be - ambiguous. Isabelle lets the Earley parser enumerate all possible - parse trees, and then tries to make the best out of the situation. - Terms that cannot be type-checked are filtered out, which often - leads to a unique result in the end. Unlike regular type - reconstruction, which is applied to the whole collection of input - terms simultaneously, the filtering stage only treats each given - term in isolation. Filtering is also not attempted for individual - types or raw ASTs (as required for @{command translations}). - - Certain warning or error messages are printed, depending on the - situation and the given configuration options. Parsing ultimately - fails, if multiple results remain after the filtering phase. - - \begin{description} - - \item @{attribute syntax_ambiguity_warning} controls output of - explicit warning messages about syntax ambiguity. - - \item @{attribute syntax_ambiguity_limit} determines the number of - resulting parse trees that are shown as part of the printed message - in case of an ambiguity. - - \end{description} -*} - - -section {* Syntax transformations \label{sec:syntax-transformations} *} - -text {* The inner syntax engine of Isabelle provides separate - mechanisms to transform parse trees either as rewrite systems on - first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs - or syntactic @{text "\"}-terms (\secref{sec:tr-funs}). This works - both for parsing and printing, as outlined in - \figref{fig:parse-print}. - - \begin{figure}[htbp] - \begin{center} - \begin{tabular}{cl} - string & \\ - @{text "\"} & lexer + parser \\ - parse tree & \\ - @{text "\"} & parse AST translation \\ - AST & \\ - @{text "\"} & AST rewriting (macros) \\ - AST & \\ - @{text "\"} & parse translation \\ - --- pre-term --- & \\ - @{text "\"} & print translation \\ - AST & \\ - @{text "\"} & AST rewriting (macros) \\ - AST & \\ - @{text "\"} & print AST translation \\ - string & - \end{tabular} - \end{center} - \caption{Parsing and printing with translations}\label{fig:parse-print} - \end{figure} - - These intermediate syntax tree formats eventually lead to a pre-term - with all names and binding scopes resolved, but most type - information still missing. Explicit type constraints might be given by - the user, or implicit position information by the system --- both - need to be passed-through carefully by syntax transformations. - - Pre-terms are further processed by the so-called \emph{check} and - \emph{unckeck} phases that are intertwined with type-inference (see - also \cite{isabelle-implementation}). The latter allows to operate - on higher-order abstract syntax with proper binding and type - information already available. - - As a rule of thumb, anything that manipulates bindings of variables - or constants needs to be implemented as syntax transformation (see - below). Anything else is better done via check/uncheck: a prominent - example application is the @{command abbreviation} concept of - Isabelle/Pure. *} - - -subsection {* Abstract syntax trees \label{sec:ast} *} - -text {* The ML datatype @{ML_type Ast.ast} explicitly represents the - intermediate AST format that is used for syntax rewriting - (\secref{sec:syn-trans}). It is defined in ML as follows: - \begin{ttbox} - datatype ast = - Constant of string | - Variable of string | - Appl of ast list - \end{ttbox} - - An AST is either an atom (constant or variable) or a list of (at - least two) subtrees. Occasional diagnostic output of ASTs uses - notation that resembles S-expression of LISP. Constant atoms are - shown as quoted strings, variable atoms as non-quoted strings and - applications as a parenthesized list of subtrees. For example, the - AST - @{ML [display] "Ast.Appl - [Ast.Constant \"_abs\", Ast.Variable \"x\", Ast.Variable \"t\"]"} - is pretty-printed as @{verbatim "(\"_abs\" x t)"}. Note that - @{verbatim "()"} and @{verbatim "(x)"} are excluded as ASTs, because - they have too few subtrees. - - \medskip AST application is merely a pro-forma mechanism to indicate - certain syntactic structures. Thus @{verbatim "(c a b)"} could mean - either term application or type application, depending on the - syntactic context. - - Nested application like @{verbatim "((\"_abs\" x t) u)"} is also - possible, but ASTs are definitely first-order: the syntax constant - @{verbatim "\"_abs\""} does not bind the @{verbatim x} in any way. - Proper bindings are introduced in later stages of the term syntax, - where @{verbatim "(\"_abs\" x t)"} becomes an @{ML Abs} node and - occurrences of @{verbatim x} in @{verbatim t} are replaced by bound - variables (represented as de-Bruijn indices). -*} - - -subsubsection {* AST constants versus variables *} - -text {* Depending on the situation --- input syntax, output syntax, - translation patterns --- the distinction of atomic asts as @{ML - Ast.Constant} versus @{ML Ast.Variable} serves slightly different - purposes. - - Input syntax of a term such as @{text "f a b = c"} does not yet - indicate the scopes of atomic entities @{text "f, a, b, c"}: they - could be global constants or local variables, even bound ones - depending on the context of the term. @{ML Ast.Variable} leaves - this choice still open: later syntax layers (or translation - functions) may capture such a variable to determine its role - specifically, to make it a constant, bound variable, free variable - etc. In contrast, syntax translations that introduce already known - constants would rather do it via @{ML Ast.Constant} to prevent - accidental re-interpretation later on. - - Output syntax turns term constants into @{ML Ast.Constant} and - variables (free or schematic) into @{ML Ast.Variable}. This - information is precise when printing fully formal @{text "\"}-terms. - - In AST translation patterns (\secref{sec:syn-trans}) the system - guesses from the current theory context which atoms should be - treated as constant versus variable for the matching process. - Sometimes this needs to be indicated more explicitly using @{text - "CONST c"} inside the term language. It is also possible to use - @{command syntax} declarations (without mixfix annotation) to - enforce that certain unqualified names are always treated as - constant within the syntax machinery. - - \medskip For ASTs that represent the language of types or sorts, the - situation is much simpler, since the concrete syntax already - distinguishes type variables from type constants (constructors). So - @{text "('a, 'b) foo"} corresponds to an AST application of some - constant for @{text foo} and variable arguments for @{text "'a"} and - @{text "'b"}. Note that the postfix application is merely a feature - of the concrete syntax, while in the AST the constructor occurs in - head position. *} - - -subsubsection {* Authentic syntax names *} - -text {* Naming constant entities within ASTs is another delicate - issue. Unqualified names are looked up in the name space tables in - the last stage of parsing, after all translations have been applied. - Since syntax transformations do not know about this later name - resolution yet, there can be surprises in boundary cases. - - \emph{Authentic syntax names} for @{ML Ast.Constant} avoid this - problem: the fully-qualified constant name with a special prefix for - its formal category (@{text "class"}, @{text "type"}, @{text - "const"}, @{text "fixed"}) represents the information faithfully - within the untyped AST format. Accidental overlap with free or - bound variables is excluded as well. Authentic syntax names work - implicitly in the following situations: - - \begin{itemize} - - \item Input of term constants (or fixed variables) that are - introduced by concrete syntax via @{command notation}: the - correspondence of a particular grammar production to some known term - entity is preserved. - - \item Input of type constants (constructors) and type classes --- - thanks to explicit syntactic distinction independently on the - context. - - \item Output of term constants, type constants, type classes --- - this information is already available from the internal term to be - printed. - - \end{itemize} - - In other words, syntax transformations that operate on input terms - written as prefix applications are difficult to make robust. - Luckily, this case rarely occurs in practice, because syntax forms - to be translated usually correspond to some bits of concrete - notation. *} - - -subsection {* Raw syntax and translations \label{sec:syn-trans} *} - -text {* - \begin{tabular}{rcll} - @{command_def "nonterminal"} & : & @{text "theory \ theory"} \\ - @{command_def "syntax"} & : & @{text "theory \ theory"} \\ - @{command_def "no_syntax"} & : & @{text "theory \ theory"} \\ - @{command_def "translations"} & : & @{text "theory \ theory"} \\ - @{command_def "no_translations"} & : & @{text "theory \ theory"} \\ - @{attribute_def syntax_ast_trace} & : & @{text attribute} & default @{text false} \\ - @{attribute_def syntax_ast_stats} & : & @{text attribute} & default @{text false} \\ - \end{tabular} - - Unlike mixfix notation for existing formal entities - (\secref{sec:notation}), raw syntax declarations provide full access - to the priority grammar of the inner syntax, without any sanity - checks. This includes additional syntactic categories (via - @{command nonterminal}) and free-form grammar productions (via - @{command syntax}). Additional syntax translations (or macros, via - @{command translations}) are required to turn resulting parse trees - into proper representations of formal entities again. - - @{rail " - @@{command nonterminal} (@{syntax name} + @'and') - ; - (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (constdecl +) - ; - (@@{command translations} | @@{command no_translations}) - (transpat ('==' | '=>' | '<=' | '\' | '\' | '\') transpat +) - ; - - constdecl: @{syntax name} '::' @{syntax type} @{syntax mixfix}? - ; - mode: ('(' ( @{syntax name} | @'output' | @{syntax name} @'output' ) ')') - ; - transpat: ('(' @{syntax nameref} ')')? @{syntax string} - "} - - \begin{description} - - \item @{command "nonterminal"}~@{text c} declares a type - constructor @{text c} (without arguments) to act as purely syntactic - type: a nonterminal symbol of the inner syntax. - - \item @{command "syntax"}~@{text "(mode) c :: \ (mx)"} augments the - priority grammar and the pretty printer table for the given print - mode (default @{verbatim "\"\""}). An optional keyword @{keyword_ref - "output"} means that only the pretty printer table is affected. - - Following \secref{sec:mixfix}, the mixfix annotation @{text "mx = - template ps q"} together with type @{text "\ = \\<^sub>1 \ \ \\<^sub>n \ \"} and - specify a grammar production. The @{text template} contains - delimiter tokens that surround @{text "n"} argument positions - (@{verbatim "_"}). The latter correspond to nonterminal symbols - @{text "A\<^sub>i"} derived from the argument types @{text "\\<^sub>i"} as - follows: - \begin{itemize} - - \item @{text "prop"} if @{text "\\<^sub>i = prop"} - - \item @{text "logic"} if @{text "\\<^sub>i = (\)\"} for logical type - constructor @{text "\ \ prop"} - - \item @{text any} if @{text "\\<^sub>i = \"} for type variables - - \item @{text "\"} if @{text "\\<^sub>i = \"} for nonterminal @{text "\"} - (syntactic type constructor) - - \end{itemize} - - Each @{text "A\<^sub>i"} is decorated by priority @{text "p\<^sub>i"} from the - given list @{text "ps"}; misssing priorities default to 0. - - The resulting nonterminal of the production is determined similarly - from type @{text "\"}, with priority @{text "q"} and default 1000. - - \medskip Parsing via this production produces parse trees @{text - "t\<^sub>1, \, t\<^sub>n"} for the argument slots. The resulting parse tree is - composed as @{text "c t\<^sub>1 \ t\<^sub>n"}, by using the syntax constant @{text - "c"} of the syntax declaration. - - Such syntactic constants are invented on the spot, without formal - check wrt.\ existing declarations. It is conventional to use plain - identifiers prefixed by a single underscore (e.g.\ @{text - "_foobar"}). Names should be chosen with care, to avoid clashes - with other syntax declarations. - - \medskip The special case of copy production is specified by @{text - "c = "}@{verbatim "\"\""} (empty string). It means that the - resulting parse tree @{text "t"} is copied directly, without any - further decoration. - - \item @{command "no_syntax"}~@{text "(mode) decls"} removes grammar - declarations (and translations) resulting from @{text decls}, which - are interpreted in the same manner as for @{command "syntax"} above. - - \item @{command "translations"}~@{text rules} specifies syntactic - translation rules (i.e.\ macros) as first-order rewrite rules on - ASTs (\secref{sec:ast}). The theory context maintains two - independent lists translation rules: parse rules (@{verbatim "=>"} - or @{text "\"}) and print rules (@{verbatim "<="} or @{text "\"}). - For convenience, both can be specified simultaneously as parse~/ - print rules (@{verbatim "=="} or @{text "\"}). - - Translation patterns may be prefixed by the syntactic category to be - used for parsing; the default is @{text logic} which means that - regular term syntax is used. Both sides of the syntax translation - rule undergo parsing and parse AST translations - \secref{sec:tr-funs}, in order to perform some fundamental - normalization like @{text "\x y. b \ \x. \y. b"}, but other AST - translation rules are \emph{not} applied recursively here. - - When processing AST patterns, the inner syntax lexer runs in a - different mode that allows identifiers to start with underscore. - This accommodates the usual naming convention for auxiliary syntax - constants --- those that do not have a logical counter part --- by - allowing to specify arbitrary AST applications within the term - syntax, independently of the corresponding concrete syntax. - - Atomic ASTs are distinguished as @{ML Ast.Constant} versus @{ML - Ast.Variable} as follows: a qualified name or syntax constant - declared via @{command syntax}, or parse tree head of concrete - notation becomes @{ML Ast.Constant}, anything else @{ML - Ast.Variable}. Note that @{text CONST} and @{text XCONST} within - the term language (\secref{sec:pure-grammar}) allow to enforce - treatment as constants. - - AST rewrite rules @{text "(lhs, rhs)"} need to obey the following - side-conditions: - - \begin{itemize} - - \item Rules must be left linear: @{text "lhs"} must not contain - repeated variables.\footnote{The deeper reason for this is that AST - equality is not well-defined: different occurrences of the ``same'' - AST could be decorated differently by accidental type-constraints or - source position information, for example.} - - \item Every variable in @{text "rhs"} must also occur in @{text - "lhs"}. - - \end{itemize} - - \item @{command "no_translations"}~@{text rules} removes syntactic - translation rules, which are interpreted in the same manner as for - @{command "translations"} above. - - \item @{attribute syntax_ast_trace} and @{attribute - syntax_ast_stats} control diagnostic output in the AST normalization - process, when translation rules are applied to concrete input or - output. - - \end{description} - - Raw syntax and translations provides a slightly more low-level - access to the grammar and the form of resulting parse trees. It is - often possible to avoid this untyped macro mechanism, and use - type-safe @{command abbreviation} or @{command notation} instead. - Some important situations where @{command syntax} and @{command - translations} are really need are as follows: - - \begin{itemize} - - \item Iterated replacement via recursive @{command translations}. - For example, consider list enumeration @{term "[a, b, c, d]"} as - defined in theory @{theory List} in Isabelle/HOL. - - \item Change of binding status of variables: anything beyond the - built-in @{keyword "binder"} mixfix annotation requires explicit - syntax translations. For example, consider list filter - comprehension @{term "[x \ xs . P]"} as defined in theory @{theory - List} in Isabelle/HOL. - - \end{itemize} -*} - -subsubsection {* Applying translation rules *} - -text {* As a term is being parsed or printed, an AST is generated as - an intermediate form according to \figref{fig:parse-print}. The AST - is normalized by applying translation rules in the manner of a - first-order term rewriting system. We first examine how a single - rule is applied. - - Let @{text "t"} be the abstract syntax tree to be normalized and - @{text "(lhs, rhs)"} some translation rule. A subtree @{text "u"} - of @{text "t"} is called \emph{redex} if it is an instance of @{text - "lhs"}; in this case the pattern @{text "lhs"} is said to match the - object @{text "u"}. A redex matched by @{text "lhs"} may be - replaced by the corresponding instance of @{text "rhs"}, thus - \emph{rewriting} the AST @{text "t"}. Matching requires some notion - of \emph{place-holders} in rule patterns: @{ML Ast.Variable} serves - this purpose. - - More precisely, the matching of the object @{text "u"} against the - pattern @{text "lhs"} is performed as follows: - - \begin{itemize} - - \item Objects of the form @{ML Ast.Variable}~@{text "x"} or @{ML - Ast.Constant}~@{text "x"} are matched by pattern @{ML - Ast.Constant}~@{text "x"}. Thus all atomic ASTs in the object are - treated as (potential) constants, and a successful match makes them - actual constants even before name space resolution (see also - \secref{sec:ast}). - - \item Object @{text "u"} is matched by pattern @{ML - Ast.Variable}~@{text "x"}, binding @{text "x"} to @{text "u"}. - - \item Object @{ML Ast.Appl}~@{text "us"} is matched by @{ML - Ast.Appl}~@{text "ts"} if @{text "us"} and @{text "ts"} have the - same length and each corresponding subtree matches. - - \item In every other case, matching fails. - - \end{itemize} - - A successful match yields a substitution that is applied to @{text - "rhs"}, generating the instance that replaces @{text "u"}. - - Normalizing an AST involves repeatedly applying translation rules - until none are applicable. This works yoyo-like: top-down, - bottom-up, top-down, etc. At each subtree position, rules are - chosen in order of appearance in the theory definitions. - - The configuration options @{attribute syntax_ast_trace} and - @{attribute syntax_ast_stats} might help to understand this process - and diagnose problems. - - \begin{warn} - If syntax translation rules work incorrectly, the output of - @{command_ref print_syntax} with its \emph{rules} sections reveals the - actual internal forms of AST pattern, without potentially confusing - concrete syntax. Recall that AST constants appear as quoted strings - and variables without quotes. - \end{warn} - - \begin{warn} - If @{attribute_ref eta_contract} is set to @{text "true"}, terms - will be @{text "\"}-contracted \emph{before} the AST rewriter sees - them. Thus some abstraction nodes needed for print rules to match - may vanish. For example, @{text "Ball A (\x. P x)"} would contract - to @{text "Ball A P"} and the standard print rule would fail to - apply. This problem can be avoided by hand-written ML translation - functions (see also \secref{sec:tr-funs}), which is in fact the same - mechanism used in built-in @{keyword "binder"} declarations. - \end{warn} -*} - - -subsection {* Syntax translation functions \label{sec:tr-funs} *} - -text {* - \begin{matharray}{rcl} - @{command_def "parse_ast_translation"} & : & @{text "theory \ theory"} \\ - @{command_def "parse_translation"} & : & @{text "theory \ theory"} \\ - @{command_def "print_translation"} & : & @{text "theory \ theory"} \\ - @{command_def "typed_print_translation"} & : & @{text "theory \ theory"} \\ - @{command_def "print_ast_translation"} & : & @{text "theory \ theory"} \\ - @{ML_antiquotation_def "class_syntax"} & : & @{text ML_antiquotation} \\ - @{ML_antiquotation_def "type_syntax"} & : & @{text ML_antiquotation} \\ - @{ML_antiquotation_def "const_syntax"} & : & @{text ML_antiquotation} \\ - @{ML_antiquotation_def "syntax_const"} & : & @{text ML_antiquotation} \\ - \end{matharray} - - Syntax translation functions written in ML admit almost arbitrary - manipulations of inner syntax, at the expense of some complexity and - obscurity in the implementation. - - @{rail " - ( @@{command parse_ast_translation} | @@{command parse_translation} | - @@{command print_translation} | @@{command typed_print_translation} | - @@{command print_ast_translation}) ('(' @'advanced' ')')? @{syntax text} - ; - (@@{ML_antiquotation class_syntax} | - @@{ML_antiquotation type_syntax} | - @@{ML_antiquotation const_syntax} | - @@{ML_antiquotation syntax_const}) name - "} - - \begin{description} - - \item @{command parse_translation} etc. declare syntax translation - functions to the theory. Any of these commands have a single - @{syntax text} argument that refers to an ML expression of - appropriate type, which are as follows by default: - - \medskip - {\footnotesize - \begin{tabular}{ll} - @{command parse_ast_translation} & : @{ML_type "(string * (Ast.ast list -> Ast.ast)) list"} \\ - @{command parse_translation} & : @{ML_type "(string * (term list -> term)) list"} \\ - @{command print_translation} & : @{ML_type "(string * (term list -> term)) list"} \\ - @{command typed_print_translation} & : @{ML_type "(string * (typ -> term list -> term)) list"} \\ - @{command print_ast_translation} & : @{ML_type "(string * (Ast.ast list -> Ast.ast)) list"} \\ - \end{tabular}} - \medskip - - The argument list consists of @{text "(c, tr)"} pairs, where @{text - "c"} is the syntax name of the formal entity involved, and @{text - "tr"} a function that translates a syntax form @{text "c args"} into - @{text "tr args"}. The ML naming convention for parse translations - is @{text "c_tr"} and for print translations @{text "c_tr'"}. - - The @{command_ref print_syntax} command displays the sets of names - associated with the translation functions of a theory under @{text - "parse_ast_translation"} etc. - - If the @{verbatim "("}@{keyword "advanced"}@{verbatim ")"} option is - given, the corresponding translation functions depend on the current - theory or proof context as additional argument. This allows to - implement advanced syntax mechanisms, as translations functions may - refer to specific theory declarations or auxiliary proof data. - - \item @{text "@{class_syntax c}"}, @{text "@{type_syntax c}"}, - @{text "@{const_syntax c}"} inline the authentic syntax name of the - given formal entities into the ML source. This is the - fully-qualified logical name prefixed by a special marker to - indicate its kind: thus different logical name spaces are properly - distinguished within parse trees. - - \item @{text "@{const_syntax c}"} inlines the name @{text "c"} of - the given syntax constant, having checked that it has been declared - via some @{command syntax} commands within the theory context. Note - that the usual naming convention makes syntax constants start with - underscore, to reduce the chance of accidental clashes with other - names occurring in parse trees (unqualified constants etc.). - - \end{description} -*} - - -subsubsection {* The translation strategy *} - -text {* The different kinds of translation functions are invoked during - the transformations between parse trees, ASTs and syntactic terms - (cf.\ \figref{fig:parse-print}). Whenever a combination of the form - @{text "c x\<^sub>1 \ x\<^sub>n"} is encountered, and a translation function - @{text "f"} of appropriate kind is declared for @{text "c"}, the - result is produced by evaluation of @{text "f [x\<^sub>1, \, x\<^sub>n]"} in ML. - - For AST translations, the arguments @{text "x\<^sub>1, \, x\<^sub>n"} are ASTs. A - combination has the form @{ML "Ast.Constant"}~@{text "c"} or @{ML - "Ast.Appl"}~@{text "["}@{ML Ast.Constant}~@{text "c, x\<^sub>1, \, x\<^sub>n]"}. - For term translations, the arguments are terms and a combination has - the form @{ML Const}~@{text "(c, \)"} or @{ML Const}~@{text "(c, \) - $ x\<^sub>1 $ \ $ x\<^sub>n"}. Terms allow more sophisticated transformations - than ASTs do, typically involving abstractions and bound - variables. \emph{Typed} print translations may even peek at the type - @{text "\"} of the constant they are invoked on, although that information - may be inaccurate. - - Regardless of whether they act on ASTs or terms, translation - functions called during the parsing process differ from those for - printing in their overall behaviour: - - \begin{description} - - \item [Parse translations] are applied bottom-up. The arguments are - already in translated form. The translations must not fail; - exceptions trigger an error message. There may be at most one - function associated with any syntactic name. - - \item [Print translations] are applied top-down. They are supplied - with arguments that are partly still in internal form. The result - again undergoes translation; therefore a print translation should - not introduce as head the very constant that invoked it. The - function may raise exception @{ML Match} to indicate failure; in - this event it has no effect. Multiple functions associated with - some syntactic name are tried in the order of declaration in the - theory. - - \end{description} - - Only constant atoms --- constructor @{ML Ast.Constant} for ASTs and - @{ML Const} for terms --- can invoke translation functions. This - means that parse translations can only be associated with parse tree - heads of concrete syntax, or syntactic constants introduced via - other translations. For plain identifiers within the term language, - the status of constant versus variable is not yet know during - parsing. This is in contrast to print translations, where constants - are explicitly known from the given term in its fully internal form. -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/ML_Tactic.thy --- a/doc-src/IsarRef/Thy/ML_Tactic.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,180 +0,0 @@ -theory ML_Tactic -imports Base Main -begin - -chapter {* ML tactic expressions *} - -text {* - Isar Proof methods closely resemble traditional tactics, when used - in unstructured sequences of @{command "apply"} commands. - Isabelle/Isar provides emulations for all major ML tactics of - classic Isabelle --- mostly for the sake of easy porting of existing - developments, as actual Isar proof texts would demand much less - diversity of proof methods. - - Unlike tactic expressions in ML, Isar proof methods provide proper - concrete syntax for additional arguments, options, modifiers etc. - Thus a typical method text is usually more concise than the - corresponding ML tactic. Furthermore, the Isar versions of classic - Isabelle tactics often cover several variant forms by a single - method with separate options to tune the behavior. For example, - method @{method simp} replaces all of @{ML simp_tac}~/ @{ML - asm_simp_tac}~/ @{ML full_simp_tac}~/ @{ML asm_full_simp_tac}, there - is also concrete syntax for augmenting the Simplifier context (the - current ``simpset'') in a convenient way. -*} - - -section {* Resolution tactics *} - -text {* - Classic Isabelle provides several variant forms of tactics for - single-step rule applications (based on higher-order resolution). - The space of resolution tactics has the following main dimensions. - - \begin{enumerate} - - \item The ``mode'' of resolution: intro, elim, destruct, or forward - (e.g.\ @{ML resolve_tac}, @{ML eresolve_tac}, @{ML dresolve_tac}, - @{ML forward_tac}). - - \item Optional explicit instantiation (e.g.\ @{ML resolve_tac} vs.\ - @{ML res_inst_tac}). - - \item Abbreviations for singleton arguments (e.g.\ @{ML resolve_tac} - vs.\ @{ML rtac}). - - \end{enumerate} - - Basically, the set of Isar tactic emulations @{method rule_tac}, - @{method erule_tac}, @{method drule_tac}, @{method frule_tac} (see - \secref{sec:tactics}) would be sufficient to cover the four modes, - either with or without instantiation, and either with single or - multiple arguments. Although it is more convenient in most cases to - use the plain @{method_ref (Pure) rule} method, or any of its - ``improper'' variants @{method erule}, @{method drule}, @{method - frule}. Note that explicit goal addressing is only supported by the - actual @{method rule_tac} version. - - With this in mind, plain resolution tactics correspond to Isar - methods as follows. - - \medskip - \begin{tabular}{lll} - @{ML rtac}~@{text "a 1"} & & @{text "rule a"} \\ - @{ML resolve_tac}~@{text "[a\<^sub>1, \] 1"} & & @{text "rule a\<^sub>1 \"} \\ - @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \] a 1"} & & - @{text "rule_tac x\<^sub>1 = t\<^sub>1 \ \ \ a"} \\[0.5ex] - @{ML rtac}~@{text "a i"} & & @{text "rule_tac [i] a"} \\ - @{ML resolve_tac}~@{text "[a\<^sub>1, \] i"} & & @{text "rule_tac [i] a\<^sub>1 \"} \\ - @{ML res_inst_tac}~@{text "ctxt [(x\<^sub>1, t\<^sub>1), \] a i"} & & - @{text "rule_tac [i] x\<^sub>1 = t\<^sub>1 \ \ \ a"} \\ - \end{tabular} - \medskip - - Note that explicit goal addressing may be usually avoided by - changing the order of subgoals with @{command "defer"} or @{command - "prefer"} (see \secref{sec:tactic-commands}). -*} - - -section {* Simplifier tactics *} - -text {* - The main Simplifier tactics @{ML simp_tac} and variants (cf.\ - \cite{isabelle-ref}) are all covered by the @{method simp} and - @{method simp_all} methods (see \secref{sec:simplifier}). Note that - there is no individual goal addressing available, simplification - acts either on the first goal (@{method simp}) or all goals - (@{method simp_all}). - - \medskip - \begin{tabular}{lll} - @{ML asm_full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp} \\ - @{ML ALLGOALS}~(@{ML asm_full_simp_tac}~@{text "@{simpset}"}) & & @{method simp_all} \\[0.5ex] - @{ML simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm)"} \\ - @{ML asm_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_simp)"} \\ - @{ML full_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(no_asm_use)"} \\ - @{ML asm_lr_simp_tac}~@{text "@{simpset} 1"} & & @{method simp}~@{text "(asm_lr)"} \\ - \end{tabular} - \medskip -*} - - -section {* Classical Reasoner tactics *} - -text {* The Classical Reasoner provides a rather large number of - variations of automated tactics, such as @{ML blast_tac}, @{ML - fast_tac}, @{ML clarify_tac} etc. The corresponding Isar methods - usually share the same base name, such as @{method blast}, @{method - fast}, @{method clarify} etc.\ (see \secref{sec:classical}). *} - - -section {* Miscellaneous tactics *} - -text {* - There are a few additional tactics defined in various theories of - Isabelle/HOL, some of these also in Isabelle/FOL or Isabelle/ZF. - The most common ones of these may be ported to Isar as follows. - - \medskip - \begin{tabular}{lll} - @{ML stac}~@{text "a 1"} & & @{text "subst a"} \\ - @{ML hyp_subst_tac}~@{text 1} & & @{text hypsubst} \\ - @{ML strip_tac}~@{text 1} & @{text "\"} & @{text "intro strip"} \\ - @{ML split_all_tac}~@{text 1} & & @{text "simp (no_asm_simp) only: split_tupled_all"} \\ - & @{text "\"} & @{text "simp only: split_tupled_all"} \\ - & @{text "\"} & @{text "clarify"} \\ - \end{tabular} -*} - - -section {* Tacticals *} - -text {* - Classic Isabelle provides a huge amount of tacticals for combination - and modification of existing tactics. This has been greatly reduced - in Isar, providing the bare minimum of combinators only: ``@{text - ","}'' (sequential composition), ``@{text "|"}'' (alternative - choices), ``@{text "?"}'' (try), ``@{text "+"}'' (repeat at least - once). These are usually sufficient in practice; if all fails, - arbitrary ML tactic code may be invoked via the @{method tactic} - method (see \secref{sec:tactics}). - - \medskip Common ML tacticals may be expressed directly in Isar as - follows: - - \medskip - \begin{tabular}{lll} - @{text "tac\<^sub>1"}~@{ML_text THEN}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1, meth\<^sub>2"} \\ - @{text "tac\<^sub>1"}~@{ML_text ORELSE}~@{text "tac\<^sub>2"} & & @{text "meth\<^sub>1 | meth\<^sub>2"} \\ - @{ML TRY}~@{text tac} & & @{text "meth?"} \\ - @{ML REPEAT1}~@{text tac} & & @{text "meth+"} \\ - @{ML REPEAT}~@{text tac} & & @{text "(meth+)?"} \\ - @{ML EVERY}~@{text "[tac\<^sub>1, \]"} & & @{text "meth\<^sub>1, \"} \\ - @{ML FIRST}~@{text "[tac\<^sub>1, \]"} & & @{text "meth\<^sub>1 | \"} \\ - \end{tabular} - \medskip - - \medskip @{ML CHANGED} (see \cite{isabelle-implementation}) is - usually not required in Isar, since most basic proof methods already - fail unless there is an actual change in the goal state. - Nevertheless, ``@{text "?"}'' (try) may be used to accept - \emph{unchanged} results as well. - - \medskip @{ML ALLGOALS}, @{ML SOMEGOAL} etc.\ (see - \cite{isabelle-implementation}) are not available in Isar, since - there is no direct goal addressing. Nevertheless, some basic - methods address all goals internally, notably @{method simp_all} - (see \secref{sec:simplifier}). Also note that @{ML ALLGOALS} can be - often replaced by ``@{text "+"}'' (repeat at least once), although - this usually has a different operational behavior: subgoals are - solved in a different order. - - \medskip Iterated resolution, such as - @{ML_text "REPEAT (FIRSTGOAL (resolve_tac ...))"}, is usually better - expressed using the @{method intro} and @{method elim} methods of - Isar (see \secref{sec:classical}). -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Misc.thy --- a/doc-src/IsarRef/Thy/Misc.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,151 +0,0 @@ -theory Misc -imports Base Main -begin - -chapter {* Other commands *} - -section {* Inspecting the context *} - -text {* - \begin{matharray}{rcl} - @{command_def "print_commands"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - @{command_def "print_theory"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_methods"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_attributes"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_theorems"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "find_theorems"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "find_consts"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "thm_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "unused_thms"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_facts"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_binds"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - \end{matharray} - - @{rail " - (@@{command print_theory} | @@{command print_theorems}) ('!'?) - ; - - @@{command find_theorems} ('(' @{syntax nat}? 'with_dups'? ')')? \\ (thmcriterion * ) - ; - thmcriterion: ('-'?) ('name' ':' @{syntax nameref} | 'intro' | 'elim' | 'dest' | - 'solves' | 'simp' ':' @{syntax term} | @{syntax term}) - ; - @@{command find_consts} (constcriterion * ) - ; - constcriterion: ('-'?) - ('name' ':' @{syntax nameref} | 'strict' ':' @{syntax type} | @{syntax type}) - ; - @@{command thm_deps} @{syntax thmrefs} - ; - @@{command unused_thms} ((@{syntax name} +) '-' (@{syntax name} * ))? - "} - - These commands print certain parts of the theory and proof context. - Note that there are some further ones available, such as for the set - of rules declared for simplifications. - - \begin{description} - - \item @{command "print_commands"} prints Isabelle's outer theory - syntax, including keywords and command. - - \item @{command "print_theory"} prints the main logical content of - the theory context; the ``@{text "!"}'' option indicates extra - verbosity. - - \item @{command "print_methods"} prints all proof methods - available in the current theory context. - - \item @{command "print_attributes"} prints all attributes - available in the current theory context. - - \item @{command "print_theorems"} prints theorems resulting from the - last command; the ``@{text "!"}'' option indicates extra verbosity. - - \item @{command "find_theorems"}~@{text criteria} retrieves facts - from the theory or proof context matching all of given search - criteria. The criterion @{text "name: p"} selects all theorems - whose fully qualified name matches pattern @{text p}, which may - contain ``@{text "*"}'' wildcards. The criteria @{text intro}, - @{text elim}, and @{text dest} select theorems that match the - current goal as introduction, elimination or destruction rules, - respectively. The criterion @{text "solves"} returns all rules - that would directly solve the current goal. The criterion - @{text "simp: t"} selects all rewrite rules whose left-hand side - matches the given term. The criterion term @{text t} selects all - theorems that contain the pattern @{text t} -- as usual, patterns - may contain occurrences of the dummy ``@{text _}'', schematic - variables, and type constraints. - - Criteria can be preceded by ``@{text "-"}'' to select theorems that - do \emph{not} match. Note that giving the empty list of criteria - yields \emph{all} currently known facts. An optional limit for the - number of printed facts may be given; the default is 40. By - default, duplicates are removed from the search result. Use - @{text with_dups} to display duplicates. - - \item @{command "find_consts"}~@{text criteria} prints all constants - whose type meets all of the given criteria. The criterion @{text - "strict: ty"} is met by any type that matches the type pattern - @{text ty}. Patterns may contain both the dummy type ``@{text _}'' - and sort constraints. The criterion @{text ty} is similar, but it - also matches against subtypes. The criterion @{text "name: p"} and - the prefix ``@{text "-"}'' function as described for @{command - "find_theorems"}. - - \item @{command "thm_deps"}~@{text "a\<^sub>1 \ a\<^sub>n"} - visualizes dependencies of facts, using Isabelle's graph browser - tool (see also \cite{isabelle-sys}). - - \item @{command "unused_thms"}~@{text "A\<^isub>1 \ A\<^isub>m - B\<^isub>1 \ B\<^isub>n"} - displays all unused theorems in theories @{text "B\<^isub>1 \ B\<^isub>n"} - or their parents, but not in @{text "A\<^isub>1 \ A\<^isub>m"} or their parents. - If @{text n} is @{text 0}, the end of the range of theories - defaults to the current theory. If no range is specified, - only the unused theorems in the current theory are displayed. - - \item @{command "print_facts"} prints all local facts of the - current context, both named and unnamed ones. - - \item @{command "print_binds"} prints all term abbreviations - present in the context. - - \end{description} -*} - - -section {* System commands *} - -text {* - \begin{matharray}{rcl} - @{command_def "cd"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - @{command_def "pwd"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - @{command_def "use_thy"}@{text "\<^sup>*"} & : & @{text "any \"} \\ - \end{matharray} - - @{rail " - (@@{command cd} | @@{command use_thy}) @{syntax name} - "} - - \begin{description} - - \item @{command "cd"}~@{text path} changes the current directory - of the Isabelle process. - - \item @{command "pwd"} prints the current working directory. - - \item @{command "use_thy"}~@{text A} preload theory @{text A}. - These system commands are scarcely used when working interactively, - since loading of theories is done automatically as required. - - \end{description} - - %FIXME proper place (!?) - Isabelle file specification may contain path variables (e.g.\ - @{verbatim "$ISABELLE_HOME"}) that are expanded accordingly. Note - that @{verbatim "~"} abbreviates @{verbatim "$USER_HOME"}, and - @{verbatim "~~"} abbreviates @{verbatim "$ISABELLE_HOME"}. The - general syntax for path specifications follows POSIX conventions. -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Outer_Syntax.thy --- a/doc-src/IsarRef/Thy/Outer_Syntax.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,427 +0,0 @@ -theory Outer_Syntax -imports Base Main -begin - -chapter {* Outer syntax --- the theory language \label{ch:outer-syntax} *} - -text {* - The rather generic framework of Isabelle/Isar syntax emerges from - three main syntactic categories: \emph{commands} of the top-level - Isar engine (covering theory and proof elements), \emph{methods} for - general goal refinements (analogous to traditional ``tactics''), and - \emph{attributes} for operations on facts (within a certain - context). Subsequently we give a reference of basic syntactic - entities underlying Isabelle/Isar syntax in a bottom-up manner. - Concrete theory and proof language elements will be introduced later - on. - - \medskip In order to get started with writing well-formed - Isabelle/Isar documents, the most important aspect to be noted is - the difference of \emph{inner} versus \emph{outer} syntax. Inner - syntax is that of Isabelle types and terms of the logic, while outer - syntax is that of Isabelle/Isar theory sources (specifications and - proofs). As a general rule, inner syntax entities may occur only as - \emph{atomic entities} within outer syntax. For example, the string - @{verbatim "\"x + y\""} and identifier @{verbatim z} are legal term - specifications within a theory, while @{verbatim "x + y"} without - quotes is not. - - Printed theory documents usually omit quotes to gain readability - (this is a matter of {\LaTeX} macro setup, say via @{verbatim - "\\isabellestyle"}, see also \cite{isabelle-sys}). Experienced - users of Isabelle/Isar may easily reconstruct the lost technical - information, while mere readers need not care about quotes at all. - - \medskip Isabelle/Isar input may contain any number of input - termination characters ``@{verbatim ";"}'' (semicolon) to separate - commands explicitly. This is particularly useful in interactive - shell sessions to make clear where the current command is intended - to end. Otherwise, the interpreter loop will continue to issue a - secondary prompt ``@{verbatim "#"}'' until an end-of-command is - clearly recognized from the input syntax, e.g.\ encounter of the - next command keyword. - - More advanced interfaces such as Proof~General \cite{proofgeneral} - do not require explicit semicolons, the amount of input text is - determined automatically by inspecting the present content of the - Emacs text buffer. In the printed presentation of Isabelle/Isar - documents semicolons are omitted altogether for readability. - - \begin{warn} - Proof~General requires certain syntax classification tables in - order to achieve properly synchronized interaction with the - Isabelle/Isar process. These tables need to be consistent with - the Isabelle version and particular logic image to be used in a - running session (common object-logics may well change the outer - syntax). The standard setup should work correctly with any of the - ``official'' logic images derived from Isabelle/HOL (including - HOLCF etc.). Users of alternative logics may need to tell - Proof~General explicitly, e.g.\ by giving an option @{verbatim "-k ZF"} - (in conjunction with @{verbatim "-l ZF"}, to specify the default - logic image). Note that option @{verbatim "-L"} does both - of this at the same time. - \end{warn} -*} - - -section {* Lexical matters \label{sec:outer-lex} *} - -text {* The outer lexical syntax consists of three main categories of - syntax tokens: - - \begin{enumerate} - - \item \emph{major keywords} --- the command names that are available - in the present logic session; - - \item \emph{minor keywords} --- additional literal tokens required - by the syntax of commands; - - \item \emph{named tokens} --- various categories of identifiers etc. - - \end{enumerate} - - Major keywords and minor keywords are guaranteed to be disjoint. - This helps user-interfaces to determine the overall structure of a - theory text, without knowing the full details of command syntax. - Internally, there is some additional information about the kind of - major keywords, which approximates the command type (theory command, - proof command etc.). - - Keywords override named tokens. For example, the presence of a - command called @{verbatim term} inhibits the identifier @{verbatim - term}, but the string @{verbatim "\"term\""} can be used instead. - By convention, the outer syntax always allows quoted strings in - addition to identifiers, wherever a named entity is expected. - - When tokenizing a given input sequence, the lexer repeatedly takes - the longest prefix of the input that forms a valid token. Spaces, - tabs, newlines and formfeeds between tokens serve as explicit - separators. - - \medskip The categories for named tokens are defined once and for - all as follows. - - \begin{center} - \begin{supertabular}{rcl} - @{syntax_def ident} & = & @{text "letter quasiletter\<^sup>*"} \\ - @{syntax_def longident} & = & @{text "ident("}@{verbatim "."}@{text "ident)\<^sup>+"} \\ - @{syntax_def symident} & = & @{text "sym\<^sup>+ | "}@{verbatim "\\"}@{verbatim "<"}@{text ident}@{verbatim ">"} \\ - @{syntax_def nat} & = & @{text "digit\<^sup>+"} \\ - @{syntax_def float} & = & @{syntax_ref nat}@{verbatim "."}@{syntax_ref nat}@{text " | "}@{verbatim "-"}@{syntax_ref nat}@{verbatim "."}@{syntax_ref nat} \\ - @{syntax_def var} & = & @{verbatim "?"}@{text "ident | "}@{verbatim "?"}@{text ident}@{verbatim "."}@{text nat} \\ - @{syntax_def typefree} & = & @{verbatim "'"}@{text ident} \\ - @{syntax_def typevar} & = & @{verbatim "?"}@{text "typefree | "}@{verbatim "?"}@{text typefree}@{verbatim "."}@{text nat} \\ - @{syntax_def string} & = & @{verbatim "\""} @{text "\"} @{verbatim "\""} \\ - @{syntax_def altstring} & = & @{verbatim "`"} @{text "\"} @{verbatim "`"} \\ - @{syntax_def verbatim} & = & @{verbatim "{*"} @{text "\"} @{verbatim "*"}@{verbatim "}"} \\[1ex] - - @{text letter} & = & @{text "latin | "}@{verbatim "\\"}@{verbatim "<"}@{text latin}@{verbatim ">"}@{text " | "}@{verbatim "\\"}@{verbatim "<"}@{text "latin latin"}@{verbatim ">"}@{text " | greek |"} \\ - & & @{verbatim "\<^isub>"}@{text " | "}@{verbatim "\<^isup>"} \\ - @{text quasiletter} & = & @{text "letter | digit | "}@{verbatim "_"}@{text " | "}@{verbatim "'"} \\ - @{text latin} & = & @{verbatim a}@{text " | \ | "}@{verbatim z}@{text " | "}@{verbatim A}@{text " | \ | "}@{verbatim Z} \\ - @{text digit} & = & @{verbatim "0"}@{text " | \ | "}@{verbatim "9"} \\ - @{text sym} & = & @{verbatim "!"}@{text " | "}@{verbatim "#"}@{text " | "}@{verbatim "$"}@{text " | "}@{verbatim "%"}@{text " | "}@{verbatim "&"}@{text " | "}@{verbatim "*"}@{text " | "}@{verbatim "+"}@{text " | "}@{verbatim "-"}@{text " | "}@{verbatim "/"}@{text " |"} \\ - & & @{verbatim "<"}@{text " | "}@{verbatim "="}@{text " | "}@{verbatim ">"}@{text " | "}@{verbatim "?"}@{text " | "}@{verbatim "@"}@{text " | "}@{verbatim "^"}@{text " | "}@{verbatim "_"}@{text " | "}@{verbatim "|"}@{text " | "}@{verbatim "~"} \\ - @{text greek} & = & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " |"} \\ - & & @{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"}@{text " | "}@{verbatim "\"} \\ - \end{supertabular} - \end{center} - - A @{syntax_ref var} or @{syntax_ref typevar} describes an unknown, - which is internally a pair of base name and index (ML type @{ML_type - indexname}). These components are either separated by a dot as in - @{text "?x.1"} or @{text "?x7.3"} or run together as in @{text - "?x1"}. The latter form is possible if the base name does not end - with digits. If the index is 0, it may be dropped altogether: - @{text "?x"} and @{text "?x0"} and @{text "?x.0"} all refer to the - same unknown, with basename @{text "x"} and index 0. - - The syntax of @{syntax_ref string} admits any characters, including - newlines; ``@{verbatim "\""}'' (double-quote) and ``@{verbatim - "\\"}'' (backslash) need to be escaped by a backslash; arbitrary - character codes may be specified as ``@{verbatim "\\"}@{text ddd}'', - with three decimal digits. Alternative strings according to - @{syntax_ref altstring} are analogous, using single back-quotes - instead. - - The body of @{syntax_ref verbatim} may consist of any text not - containing ``@{verbatim "*"}@{verbatim "}"}''; this allows - convenient inclusion of quotes without further escapes. There is no - way to escape ``@{verbatim "*"}@{verbatim "}"}''. If the quoted - text is {\LaTeX} source, one may usually add some blank or comment - to avoid the critical character sequence. - - Source comments take the form @{verbatim "(*"}~@{text - "\"}~@{verbatim "*)"} and may be nested, although the user-interface - might prevent this. Note that this form indicates source comments - only, which are stripped after lexical analysis of the input. The - Isar syntax also provides proper \emph{document comments} that are - considered as part of the text (see \secref{sec:comments}). - - Common mathematical symbols such as @{text \} are represented in - Isabelle as @{verbatim \}. There are infinitely many Isabelle - symbols like this, although proper presentation is left to front-end - tools such as {\LaTeX}, Proof~General, or Isabelle/jEdit. A list of - predefined Isabelle symbols that work well with these tools is given - in \appref{app:symbols}. Note that @{verbatim "\"} does not belong - to the @{text letter} category, since it is already used differently - in the Pure term language. *} - - -section {* Common syntax entities *} - -text {* - We now introduce several basic syntactic entities, such as names, - terms, and theorem specifications, which are factored out of the - actual Isar language elements to be described later. -*} - - -subsection {* Names *} - -text {* Entity @{syntax name} usually refers to any name of types, - constants, theorems etc.\ that are to be \emph{declared} or - \emph{defined} (so qualified identifiers are excluded here). Quoted - strings provide an escape for non-identifier names or those ruled - out by outer syntax keywords (e.g.\ quoted @{verbatim "\"let\""}). - Already existing objects are usually referenced by @{syntax - nameref}. - - @{rail " - @{syntax_def name}: @{syntax ident} | @{syntax symident} | - @{syntax string} | @{syntax nat} - ; - @{syntax_def parname}: '(' @{syntax name} ')' - ; - @{syntax_def nameref}: @{syntax name} | @{syntax longident} - "} -*} - - -subsection {* Numbers *} - -text {* The outer lexical syntax (\secref{sec:outer-lex}) admits - natural numbers and floating point numbers. These are combined as - @{syntax int} and @{syntax real} as follows. - - @{rail " - @{syntax_def int}: @{syntax nat} | '-' @{syntax nat} - ; - @{syntax_def real}: @{syntax float} | @{syntax int} - "} - - Note that there is an overlap with the category @{syntax name}, - which also includes @{syntax nat}. -*} - - -subsection {* Comments \label{sec:comments} *} - -text {* Large chunks of plain @{syntax text} are usually given - @{syntax verbatim}, i.e.\ enclosed in @{verbatim "{"}@{verbatim - "*"}~@{text "\"}~@{verbatim "*"}@{verbatim "}"}. For convenience, - any of the smaller text units conforming to @{syntax nameref} are - admitted as well. A marginal @{syntax comment} is of the form - @{verbatim "--"}~@{syntax text}. Any number of these may occur - within Isabelle/Isar commands. - - @{rail " - @{syntax_def text}: @{syntax verbatim} | @{syntax nameref} - ; - @{syntax_def comment}: '--' @{syntax text} - "} -*} - - -subsection {* Type classes, sorts and arities *} - -text {* - Classes are specified by plain names. Sorts have a very simple - inner syntax, which is either a single class name @{text c} or a - list @{text "{c\<^sub>1, \, c\<^sub>n}"} referring to the - intersection of these classes. The syntax of type arities is given - directly at the outer level. - - @{rail " - @{syntax_def classdecl}: @{syntax name} (('<' | '\') (@{syntax nameref} + ','))? - ; - @{syntax_def sort}: @{syntax nameref} - ; - @{syntax_def arity}: ('(' (@{syntax sort} + ',') ')')? @{syntax sort} - "} -*} - - -subsection {* Types and terms \label{sec:types-terms} *} - -text {* - The actual inner Isabelle syntax, that of types and terms of the - logic, is far too sophisticated in order to be modelled explicitly - at the outer theory level. Basically, any such entity has to be - quoted to turn it into a single token (the parsing and type-checking - is performed internally later). For convenience, a slightly more - liberal convention is adopted: quotes may be omitted for any type or - term that is already atomic at the outer level. For example, one - may just write @{verbatim x} instead of quoted @{verbatim "\"x\""}. - Note that symbolic identifiers (e.g.\ @{verbatim "++"} or @{text - "\"} are available as well, provided these have not been superseded - by commands or other keywords already (such as @{verbatim "="} or - @{verbatim "+"}). - - @{rail " - @{syntax_def type}: @{syntax nameref} | @{syntax typefree} | - @{syntax typevar} - ; - @{syntax_def term}: @{syntax nameref} | @{syntax var} - ; - @{syntax_def prop}: @{syntax term} - "} - - Positional instantiations are indicated by giving a sequence of - terms, or the placeholder ``@{text _}'' (underscore), which means to - skip a position. - - @{rail " - @{syntax_def inst}: '_' | @{syntax term} - ; - @{syntax_def insts}: (@{syntax inst} *) - "} - - Type declarations and definitions usually refer to @{syntax - typespec} on the left-hand side. This models basic type constructor - application at the outer syntax level. Note that only plain postfix - notation is available here, but no infixes. - - @{rail " - @{syntax_def typespec}: - (() | @{syntax typefree} | '(' ( @{syntax typefree} + ',' ) ')') @{syntax name} - ; - @{syntax_def typespec_sorts}: - (() | (@{syntax typefree} ('::' @{syntax sort})?) | - '(' ( (@{syntax typefree} ('::' @{syntax sort})?) + ',' ) ')') @{syntax name} - "} -*} - - -subsection {* Term patterns and declarations \label{sec:term-decls} *} - -text {* Wherever explicit propositions (or term fragments) occur in a - proof text, casual binding of schematic term variables may be given - specified via patterns of the form ``@{text "(\ p\<^sub>1 \ p\<^sub>n)"}''. - This works both for @{syntax term} and @{syntax prop}. - - @{rail " - @{syntax_def term_pat}: '(' (@'is' @{syntax term} +) ')' - ; - @{syntax_def prop_pat}: '(' (@'is' @{syntax prop} +) ')' - "} - - \medskip Declarations of local variables @{text "x :: \"} and - logical propositions @{text "a : \"} represent different views on - the same principle of introducing a local scope. In practice, one - may usually omit the typing of @{syntax vars} (due to - type-inference), and the naming of propositions (due to implicit - references of current facts). In any case, Isar proof elements - usually admit to introduce multiple such items simultaneously. - - @{rail " - @{syntax_def vars}: (@{syntax name} +) ('::' @{syntax type})? - ; - @{syntax_def props}: @{syntax thmdecl}? (@{syntax prop} @{syntax prop_pat}? +) - "} - - The treatment of multiple declarations corresponds to the - complementary focus of @{syntax vars} versus @{syntax props}. In - ``@{text "x\<^sub>1 \ x\<^sub>n :: \"}'' the typing refers to all variables, while - in @{text "a: \\<^sub>1 \ \\<^sub>n"} the naming refers to all propositions - collectively. Isar language elements that refer to @{syntax vars} - or @{syntax props} typically admit separate typings or namings via - another level of iteration, with explicit @{keyword_ref "and"} - separators; e.g.\ see @{command "fix"} and @{command "assume"} in - \secref{sec:proof-context}. -*} - - -subsection {* Attributes and theorems \label{sec:syn-att} *} - -text {* Attributes have their own ``semi-inner'' syntax, in the sense - that input conforming to @{syntax args} below is parsed by the - attribute a second time. The attribute argument specifications may - be any sequence of atomic entities (identifiers, strings etc.), or - properly bracketed argument lists. Below @{syntax atom} refers to - any atomic entity, including any @{syntax keyword} conforming to - @{syntax symident}. - - @{rail " - @{syntax_def atom}: @{syntax nameref} | @{syntax typefree} | - @{syntax typevar} | @{syntax var} | @{syntax nat} | @{syntax float} | - @{syntax keyword} - ; - arg: @{syntax atom} | '(' @{syntax args} ')' | '[' @{syntax args} ']' - ; - @{syntax_def args}: arg * - ; - @{syntax_def attributes}: '[' (@{syntax nameref} @{syntax args} * ',') ']' - "} - - Theorem specifications come in several flavors: @{syntax axmdecl} - and @{syntax thmdecl} usually refer to axioms, assumptions or - results of goal statements, while @{syntax thmdef} collects lists of - existing theorems. Existing theorems are given by @{syntax thmref} - and @{syntax thmrefs}, the former requires an actual singleton - result. - - There are three forms of theorem references: - \begin{enumerate} - - \item named facts @{text "a"}, - - \item selections from named facts @{text "a(i)"} or @{text "a(j - k)"}, - - \item literal fact propositions using @{syntax_ref altstring} syntax - @{verbatim "`"}@{text "\"}@{verbatim "`"} (see also method - @{method_ref fact}). - - \end{enumerate} - - Any kind of theorem specification may include lists of attributes - both on the left and right hand sides; attributes are applied to any - immediately preceding fact. If names are omitted, the theorems are - not stored within the theorem database of the theory or proof - context, but any given attributes are applied nonetheless. - - An extra pair of brackets around attributes (like ``@{text - "[[simproc a]]"}'') abbreviates a theorem reference involving an - internal dummy fact, which will be ignored later on. So only the - effect of the attribute on the background context will persist. - This form of in-place declarations is particularly useful with - commands like @{command "declare"} and @{command "using"}. - - @{rail " - @{syntax_def axmdecl}: @{syntax name} @{syntax attributes}? ':' - ; - @{syntax_def thmdecl}: thmbind ':' - ; - @{syntax_def thmdef}: thmbind '=' - ; - @{syntax_def thmref}: - (@{syntax nameref} selection? | @{syntax altstring}) @{syntax attributes}? | - '[' @{syntax attributes} ']' - ; - @{syntax_def thmrefs}: @{syntax thmref} + - ; - - thmbind: @{syntax name} @{syntax attributes} | @{syntax name} | @{syntax attributes} - ; - selection: '(' ((@{syntax nat} | @{syntax nat} '-' @{syntax nat}?) + ',') ')' - "} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Preface.thy --- a/doc-src/IsarRef/Thy/Preface.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,70 +0,0 @@ -theory Preface -imports Base Main -begin - -chapter {* Preface *} - -text {* - The \emph{Isabelle} system essentially provides a generic - infrastructure for building deductive systems (programmed in - Standard ML), with a special focus on interactive theorem proving in - higher-order logics. Many years ago, even end-users would refer to - certain ML functions (goal commands, tactics, tacticals etc.) to - pursue their everyday theorem proving tasks. - - In contrast \emph{Isar} provides an interpreted language environment - of its own, which has been specifically tailored for the needs of - theory and proof development. Compared to raw ML, the Isabelle/Isar - top-level provides a more robust and comfortable development - platform, with proper support for theory development graphs, managed - transactions with unlimited undo etc. The Isabelle/Isar version of - the \emph{Proof~General} user interface - \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end - for interactive theory and proof development in this advanced - theorem proving environment, even though it is somewhat biased - towards old-style proof scripts. - - \medskip Apart from the technical advances over bare-bones ML - programming, the main purpose of the Isar language is to provide a - conceptually different view on machine-checked proofs - \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for - \emph{Intelligible semi-automated reasoning}. Drawing from both the - traditions of informal mathematical proof texts and high-level - programming languages, Isar offers a versatile environment for - structured formal proof documents. Thus properly written Isar - proofs become accessible to a broader audience than unstructured - tactic scripts (which typically only provide operational information - for the machine). Writing human-readable proof texts certainly - requires some additional efforts by the writer to achieve a good - presentation, both of formal and informal parts of the text. On the - other hand, human-readable formal texts gain some value in their own - right, independently of the mechanic proof-checking process. - - Despite its grand design of structured proof texts, Isar is able to - assimilate the old tactical style as an ``improper'' sub-language. - This provides an easy upgrade path for existing tactic scripts, as - well as some means for interactive experimentation and debugging of - structured proofs. Isabelle/Isar supports a broad range of proof - styles, both readable and unreadable ones. - - \medskip The generic Isabelle/Isar framework (see - \chref{ch:isar-framework}) works reasonably well for any Isabelle - object-logic that conforms to the natural deduction view of the - Isabelle/Pure framework. Specific language elements introduced by - Isabelle/HOL are described in \chref{ch:hol}. Although the main - language elements are already provided by the Isabelle/Pure - framework, examples given in the generic parts will usually refer to - Isabelle/HOL. - - \medskip Isar commands may be either \emph{proper} document - constructors, or \emph{improper commands}. Some proof methods and - attributes introduced later are classified as improper as well. - Improper Isar language elements, which are marked by ``@{text - "\<^sup>*"}'' in the subsequent chapters; they are often helpful - when developing proof documents, but their use is discouraged for - the final human-readable outcome. Typical examples are diagnostic - commands that print terms or theorems according to the current - context; other commands emulate old-style tactical theorem proving. -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Proof.thy --- a/doc-src/IsarRef/Thy/Proof.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1522 +0,0 @@ -theory Proof -imports Base Main -begin - -chapter {* Proofs \label{ch:proofs} *} - -text {* - Proof commands perform transitions of Isar/VM machine - configurations, which are block-structured, consisting of a stack of - nodes with three main components: logical proof context, current - facts, and open goals. Isar/VM transitions are typed according to - the following three different modes of operation: - - \begin{description} - - \item @{text "proof(prove)"} means that a new goal has just been - stated that is now to be \emph{proven}; the next command may refine - it by some proof method, and enter a sub-proof to establish the - actual result. - - \item @{text "proof(state)"} is like a nested theory mode: the - context may be augmented by \emph{stating} additional assumptions, - intermediate results etc. - - \item @{text "proof(chain)"} is intermediate between @{text - "proof(state)"} and @{text "proof(prove)"}: existing facts (i.e.\ - the contents of the special ``@{fact_ref this}'' register) have been - just picked up in order to be used when refining the goal claimed - next. - - \end{description} - - The proof mode indicator may be understood as an instruction to the - writer, telling what kind of operation may be performed next. The - corresponding typings of proof commands restricts the shape of - well-formed proof texts to particular command sequences. So dynamic - arrangements of commands eventually turn out as static texts of a - certain structure. - - \Appref{ap:refcard} gives a simplified grammar of the (extensible) - language emerging that way from the different types of proof - commands. The main ideas of the overall Isar framework are - explained in \chref{ch:isar-framework}. -*} - - -section {* Proof structure *} - -subsection {* Formal notepad *} - -text {* - \begin{matharray}{rcl} - @{command_def "notepad"} & : & @{text "local_theory \ proof(state)"} \\ - \end{matharray} - - @{rail " - @@{command notepad} @'begin' - ; - @@{command end} - "} - - \begin{description} - - \item @{command "notepad"}~@{keyword "begin"} opens a proof state - without any goal statement. This allows to experiment with Isar, - without producing any persistent result. - - The notepad can be closed by @{command "end"} or discontinued by - @{command "oops"}. - - \end{description} -*} - - -subsection {* Blocks *} - -text {* - \begin{matharray}{rcl} - @{command_def "next"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "{"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "}"} & : & @{text "proof(state) \ proof(state)"} \\ - \end{matharray} - - While Isar is inherently block-structured, opening and closing - blocks is mostly handled rather casually, with little explicit - user-intervention. Any local goal statement automatically opens - \emph{two} internal blocks, which are closed again when concluding - the sub-proof (by @{command "qed"} etc.). Sections of different - context within a sub-proof may be switched via @{command "next"}, - which is just a single block-close followed by block-open again. - The effect of @{command "next"} is to reset the local proof context; - there is no goal focus involved here! - - For slightly more advanced applications, there are explicit block - parentheses as well. These typically achieve a stronger forward - style of reasoning. - - \begin{description} - - \item @{command "next"} switches to a fresh block within a - sub-proof, resetting the local context to the initial one. - - \item @{command "{"} and @{command "}"} explicitly open and close - blocks. Any current facts pass through ``@{command "{"}'' - unchanged, while ``@{command "}"}'' causes any result to be - \emph{exported} into the enclosing context. Thus fixed variables - are generalized, assumptions discharged, and local definitions - unfolded (cf.\ \secref{sec:proof-context}). There is no difference - of @{command "assume"} and @{command "presume"} in this mode of - forward reasoning --- in contrast to plain backward reasoning with - the result exported at @{command "show"} time. - - \end{description} -*} - - -subsection {* Omitting proofs *} - -text {* - \begin{matharray}{rcl} - @{command_def "oops"} & : & @{text "proof \ local_theory | theory"} \\ - \end{matharray} - - The @{command "oops"} command discontinues the current proof - attempt, while considering the partial proof text as properly - processed. This is conceptually quite different from ``faking'' - actual proofs via @{command_ref "sorry"} (see - \secref{sec:proof-steps}): @{command "oops"} does not observe the - proof structure at all, but goes back right to the theory level. - Furthermore, @{command "oops"} does not produce any result theorem - --- there is no intended claim to be able to complete the proof - in any way. - - A typical application of @{command "oops"} is to explain Isar proofs - \emph{within} the system itself, in conjunction with the document - preparation tools of Isabelle described in \chref{ch:document-prep}. - Thus partial or even wrong proof attempts can be discussed in a - logically sound manner. Note that the Isabelle {\LaTeX} macros can - be easily adapted to print something like ``@{text "\"}'' instead of - the keyword ``@{command "oops"}''. -*} - - -section {* Statements *} - -subsection {* Context elements \label{sec:proof-context} *} - -text {* - \begin{matharray}{rcl} - @{command_def "fix"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "assume"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "presume"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "def"} & : & @{text "proof(state) \ proof(state)"} \\ - \end{matharray} - - The logical proof context consists of fixed variables and - assumptions. The former closely correspond to Skolem constants, or - meta-level universal quantification as provided by the Isabelle/Pure - logical framework. Introducing some \emph{arbitrary, but fixed} - variable via ``@{command "fix"}~@{text x}'' results in a local value - that may be used in the subsequent proof as any other variable or - constant. Furthermore, any result @{text "\ \[x]"} exported from - the context will be universally closed wrt.\ @{text x} at the - outermost level: @{text "\ \x. \[x]"} (this is expressed in normal - form using Isabelle's meta-variables). - - Similarly, introducing some assumption @{text \} has two effects. - On the one hand, a local theorem is created that may be used as a - fact in subsequent proof steps. On the other hand, any result - @{text "\ \ \"} exported from the context becomes conditional wrt.\ - the assumption: @{text "\ \ \ \"}. Thus, solving an enclosing goal - using such a result would basically introduce a new subgoal stemming - from the assumption. How this situation is handled depends on the - version of assumption command used: while @{command "assume"} - insists on solving the subgoal by unification with some premise of - the goal, @{command "presume"} leaves the subgoal unchanged in order - to be proved later by the user. - - Local definitions, introduced by ``@{command "def"}~@{text "x \ - t"}'', are achieved by combining ``@{command "fix"}~@{text x}'' with - another version of assumption that causes any hypothetical equation - @{text "x \ t"} to be eliminated by the reflexivity rule. Thus, - exporting some result @{text "x \ t \ \[x]"} yields @{text "\ - \[t]"}. - - @{rail " - @@{command fix} (@{syntax vars} + @'and') - ; - (@@{command assume} | @@{command presume}) (@{syntax props} + @'and') - ; - @@{command def} (def + @'and') - ; - def: @{syntax thmdecl}? \\ @{syntax name} ('==' | '\') @{syntax term} @{syntax term_pat}? - "} - - \begin{description} - - \item @{command "fix"}~@{text x} introduces a local variable @{text - x} that is \emph{arbitrary, but fixed.} - - \item @{command "assume"}~@{text "a: \"} and @{command - "presume"}~@{text "a: \"} introduce a local fact @{text "\ \ \"} by - assumption. Subsequent results applied to an enclosing goal (e.g.\ - by @{command_ref "show"}) are handled as follows: @{command - "assume"} expects to be able to unify with existing premises in the - goal, while @{command "presume"} leaves @{text \} as new subgoals. - - Several lists of assumptions may be given (separated by - @{keyword_ref "and"}; the resulting list of current facts consists - of all of these concatenated. - - \item @{command "def"}~@{text "x \ t"} introduces a local - (non-polymorphic) definition. In results exported from the context, - @{text x} is replaced by @{text t}. Basically, ``@{command - "def"}~@{text "x \ t"}'' abbreviates ``@{command "fix"}~@{text - x}~@{command "assume"}~@{text "x \ t"}'', with the resulting - hypothetical equation solved by reflexivity. - - The default name for the definitional equation is @{text x_def}. - Several simultaneous definitions may be given at the same time. - - \end{description} - - The special name @{fact_ref prems} refers to all assumptions of the - current context as a list of theorems. This feature should be used - with great care! It is better avoided in final proof texts. -*} - - -subsection {* Term abbreviations \label{sec:term-abbrev} *} - -text {* - \begin{matharray}{rcl} - @{command_def "let"} & : & @{text "proof(state) \ proof(state)"} \\ - @{keyword_def "is"} & : & syntax \\ - \end{matharray} - - Abbreviations may be either bound by explicit @{command - "let"}~@{text "p \ t"} statements, or by annotating assumptions or - goal statements with a list of patterns ``@{text "(\ p\<^sub>1 \ - p\<^sub>n)"}''. In both cases, higher-order matching is invoked to - bind extra-logical term variables, which may be either named - schematic variables of the form @{text ?x}, or nameless dummies - ``@{variable _}'' (underscore). Note that in the @{command "let"} - form the patterns occur on the left-hand side, while the @{keyword - "is"} patterns are in postfix position. - - Polymorphism of term bindings is handled in Hindley-Milner style, - similar to ML. Type variables referring to local assumptions or - open goal statements are \emph{fixed}, while those of finished - results or bound by @{command "let"} may occur in \emph{arbitrary} - instances later. Even though actual polymorphism should be rarely - used in practice, this mechanism is essential to achieve proper - incremental type-inference, as the user proceeds to build up the - Isar proof text from left to right. - - \medskip Term abbreviations are quite different from local - definitions as introduced via @{command "def"} (see - \secref{sec:proof-context}). The latter are visible within the - logic as actual equations, while abbreviations disappear during the - input process just after type checking. Also note that @{command - "def"} does not support polymorphism. - - @{rail " - @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and') - "} - - The syntax of @{keyword "is"} patterns follows @{syntax term_pat} or - @{syntax prop_pat} (see \secref{sec:term-decls}). - - \begin{description} - - \item @{command "let"}~@{text "p\<^sub>1 = t\<^sub>1 \ \ p\<^sub>n = t\<^sub>n"} binds any - text variables in patterns @{text "p\<^sub>1, \, p\<^sub>n"} by simultaneous - higher-order matching against terms @{text "t\<^sub>1, \, t\<^sub>n"}. - - \item @{text "(\ p\<^sub>1 \ p\<^sub>n)"} resembles @{command "let"}, but - matches @{text "p\<^sub>1, \, p\<^sub>n"} against the preceding statement. Also - note that @{keyword "is"} is not a separate command, but part of - others (such as @{command "assume"}, @{command "have"} etc.). - - \end{description} - - Some \emph{implicit} term abbreviations\index{term abbreviations} - for goals and facts are available as well. For any open goal, - @{variable_ref thesis} refers to its object-level statement, - abstracted over any meta-level parameters (if present). Likewise, - @{variable_ref this} is bound for fact statements resulting from - assumptions or finished goals. In case @{variable this} refers to - an object-logic statement that is an application @{text "f t"}, then - @{text t} is bound to the special text variable ``@{variable "\"}'' - (three dots). The canonical application of this convenience are - calculational proofs (see \secref{sec:calculation}). -*} - - -subsection {* Facts and forward chaining \label{sec:proof-facts} *} - -text {* - \begin{matharray}{rcl} - @{command_def "note"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "then"} & : & @{text "proof(state) \ proof(chain)"} \\ - @{command_def "from"} & : & @{text "proof(state) \ proof(chain)"} \\ - @{command_def "with"} & : & @{text "proof(state) \ proof(chain)"} \\ - @{command_def "using"} & : & @{text "proof(prove) \ proof(prove)"} \\ - @{command_def "unfolding"} & : & @{text "proof(prove) \ proof(prove)"} \\ - \end{matharray} - - New facts are established either by assumption or proof of local - statements. Any fact will usually be involved in further proofs, - either as explicit arguments of proof methods, or when forward - chaining towards the next goal via @{command "then"} (and variants); - @{command "from"} and @{command "with"} are composite forms - involving @{command "note"}. The @{command "using"} elements - augments the collection of used facts \emph{after} a goal has been - stated. Note that the special theorem name @{fact_ref this} refers - to the most recently established facts, but only \emph{before} - issuing a follow-up claim. - - @{rail " - @@{command note} (@{syntax thmdef}? @{syntax thmrefs} + @'and') - ; - (@@{command from} | @@{command with} | @@{command using} | @@{command unfolding}) - (@{syntax thmrefs} + @'and') - "} - - \begin{description} - - \item @{command "note"}~@{text "a = b\<^sub>1 \ b\<^sub>n"} recalls existing facts - @{text "b\<^sub>1, \, b\<^sub>n"}, binding the result as @{text a}. Note that - attributes may be involved as well, both on the left and right hand - sides. - - \item @{command "then"} indicates forward chaining by the current - facts in order to establish the goal to be claimed next. The - initial proof method invoked to refine that will be offered the - facts to do ``anything appropriate'' (see also - \secref{sec:proof-steps}). For example, method @{method (Pure) rule} - (see \secref{sec:pure-meth-att}) would typically do an elimination - rather than an introduction. Automatic methods usually insert the - facts into the goal state before operation. This provides a simple - scheme to control relevance of facts in automated proof search. - - \item @{command "from"}~@{text b} abbreviates ``@{command - "note"}~@{text b}~@{command "then"}''; thus @{command "then"} is - equivalent to ``@{command "from"}~@{text this}''. - - \item @{command "with"}~@{text "b\<^sub>1 \ b\<^sub>n"} abbreviates ``@{command - "from"}~@{text "b\<^sub>1 \ b\<^sub>n \ this"}''; thus the forward chaining - is from earlier facts together with the current ones. - - \item @{command "using"}~@{text "b\<^sub>1 \ b\<^sub>n"} augments the facts being - currently indicated for use by a subsequent refinement step (such as - @{command_ref "apply"} or @{command_ref "proof"}). - - \item @{command "unfolding"}~@{text "b\<^sub>1 \ b\<^sub>n"} is structurally - similar to @{command "using"}, but unfolds definitional equations - @{text "b\<^sub>1, \ b\<^sub>n"} throughout the goal state and facts. - - \end{description} - - Forward chaining with an empty list of theorems is the same as not - chaining at all. Thus ``@{command "from"}~@{text nothing}'' has no - effect apart from entering @{text "prove(chain)"} mode, since - @{fact_ref nothing} is bound to the empty list of theorems. - - Basic proof methods (such as @{method_ref (Pure) rule}) expect multiple - facts to be given in their proper order, corresponding to a prefix - of the premises of the rule involved. Note that positions may be - easily skipped using something like @{command "from"}~@{text "_ - \ a \ b"}, for example. This involves the trivial rule - @{text "PROP \ \ PROP \"}, which is bound in Isabelle/Pure as - ``@{fact_ref "_"}'' (underscore). - - Automated methods (such as @{method simp} or @{method auto}) just - insert any given facts before their usual operation. Depending on - the kind of procedure involved, the order of facts is less - significant here. -*} - - -subsection {* Goals \label{sec:goals} *} - -text {* - \begin{matharray}{rcl} - @{command_def "lemma"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "theorem"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "corollary"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "schematic_lemma"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "schematic_theorem"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "schematic_corollary"} & : & @{text "local_theory \ proof(prove)"} \\ - @{command_def "have"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ - @{command_def "show"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ - @{command_def "hence"} & : & @{text "proof(state) \ proof(prove)"} \\ - @{command_def "thus"} & : & @{text "proof(state) \ proof(prove)"} \\ - @{command_def "print_statement"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - \end{matharray} - - From a theory context, proof mode is entered by an initial goal - command such as @{command "lemma"}, @{command "theorem"}, or - @{command "corollary"}. Within a proof, new claims may be - introduced locally as well; four variants are available here to - indicate whether forward chaining of facts should be performed - initially (via @{command_ref "then"}), and whether the final result - is meant to solve some pending goal. - - Goals may consist of multiple statements, resulting in a list of - facts eventually. A pending multi-goal is internally represented as - a meta-level conjunction (@{text "&&&"}), which is usually - split into the corresponding number of sub-goals prior to an initial - method application, via @{command_ref "proof"} - (\secref{sec:proof-steps}) or @{command_ref "apply"} - (\secref{sec:tactic-commands}). The @{method_ref induct} method - covered in \secref{sec:cases-induct} acts on multiple claims - simultaneously. - - Claims at the theory level may be either in short or long form. A - short goal merely consists of several simultaneous propositions - (often just one). A long goal includes an explicit context - specification for the subsequent conclusion, involving local - parameters and assumptions. Here the role of each part of the - statement is explicitly marked by separate keywords (see also - \secref{sec:locale}); the local assumptions being introduced here - are available as @{fact_ref assms} in the proof. Moreover, there - are two kinds of conclusions: @{element_def "shows"} states several - simultaneous propositions (essentially a big conjunction), while - @{element_def "obtains"} claims several simultaneous simultaneous - contexts of (essentially a big disjunction of eliminated parameters - and assumptions, cf.\ \secref{sec:obtain}). - - @{rail " - (@@{command lemma} | @@{command theorem} | @@{command corollary} | - @@{command schematic_lemma} | @@{command schematic_theorem} | - @@{command schematic_corollary}) @{syntax target}? (goal | longgoal) - ; - (@@{command have} | @@{command show} | @@{command hence} | @@{command thus}) goal - ; - @@{command print_statement} @{syntax modes}? @{syntax thmrefs} - ; - - goal: (@{syntax props} + @'and') - ; - longgoal: @{syntax thmdecl}? (@{syntax_ref \"includes\"}?) (@{syntax context_elem} * ) conclusion - ; - conclusion: @'shows' goal | @'obtains' (@{syntax parname}? case + '|') - ; - case: (@{syntax vars} + @'and') @'where' (@{syntax props} + @'and') - "} - - \begin{description} - - \item @{command "lemma"}~@{text "a: \"} enters proof mode with - @{text \} as main goal, eventually resulting in some fact @{text "\ - \"} to be put back into the target context. An additional @{syntax - context} specification may build up an initial proof context for the - subsequent claim; this includes local definitions and syntax as - well, see also @{syntax "includes"} in \secref{sec:bundle} and - @{syntax context_elem} in \secref{sec:locale}. - - \item @{command "theorem"}~@{text "a: \"} and @{command - "corollary"}~@{text "a: \"} are essentially the same as @{command - "lemma"}~@{text "a: \"}, but the facts are internally marked as - being of a different kind. This discrimination acts like a formal - comment. - - \item @{command "schematic_lemma"}, @{command "schematic_theorem"}, - @{command "schematic_corollary"} are similar to @{command "lemma"}, - @{command "theorem"}, @{command "corollary"}, respectively but allow - the statement to contain unbound schematic variables. - - Under normal circumstances, an Isar proof text needs to specify - claims explicitly. Schematic goals are more like goals in Prolog, - where certain results are synthesized in the course of reasoning. - With schematic statements, the inherent compositionality of Isar - proofs is lost, which also impacts performance, because proof - checking is forced into sequential mode. - - \item @{command "have"}~@{text "a: \"} claims a local goal, - eventually resulting in a fact within the current logical context. - This operation is completely independent of any pending sub-goals of - an enclosing goal statements, so @{command "have"} may be freely - used for experimental exploration of potential results within a - proof body. - - \item @{command "show"}~@{text "a: \"} is like @{command - "have"}~@{text "a: \"} plus a second stage to refine some pending - sub-goal for each one of the finished result, after having been - exported into the corresponding context (at the head of the - sub-proof of this @{command "show"} command). - - To accommodate interactive debugging, resulting rules are printed - before being applied internally. Even more, interactive execution - of @{command "show"} predicts potential failure and displays the - resulting error as a warning beforehand. Watch out for the - following message: - - %FIXME proper antiquitation - \begin{ttbox} - Problem! Local statement will fail to solve any pending goal - \end{ttbox} - - \item @{command "hence"} abbreviates ``@{command "then"}~@{command - "have"}'', i.e.\ claims a local goal to be proven by forward - chaining the current facts. Note that @{command "hence"} is also - equivalent to ``@{command "from"}~@{text this}~@{command "have"}''. - - \item @{command "thus"} abbreviates ``@{command "then"}~@{command - "show"}''. Note that @{command "thus"} is also equivalent to - ``@{command "from"}~@{text this}~@{command "show"}''. - - \item @{command "print_statement"}~@{text a} prints facts from the - current theory or proof context in long statement form, according to - the syntax for @{command "lemma"} given above. - - \end{description} - - Any goal statement causes some term abbreviations (such as - @{variable_ref "?thesis"}) to be bound automatically, see also - \secref{sec:term-abbrev}. - - The optional case names of @{element_ref "obtains"} have a twofold - meaning: (1) during the of this claim they refer to the the local - context introductions, (2) the resulting rule is annotated - accordingly to support symbolic case splits when used with the - @{method_ref cases} method (cf.\ \secref{sec:cases-induct}). -*} - - -section {* Refinement steps *} - -subsection {* Proof method expressions \label{sec:proof-meth} *} - -text {* Proof methods are either basic ones, or expressions composed - of methods via ``@{verbatim ","}'' (sequential composition), - ``@{verbatim "|"}'' (alternative choices), ``@{verbatim "?"}'' - (try), ``@{verbatim "+"}'' (repeat at least once), ``@{verbatim - "["}@{text n}@{verbatim "]"}'' (restriction to first @{text n} - sub-goals, with default @{text "n = 1"}). In practice, proof - methods are usually just a comma separated list of @{syntax - nameref}~@{syntax args} specifications. Note that parentheses may - be dropped for single method specifications (with no arguments). - - @{rail " - @{syntax_def method}: - (@{syntax nameref} | '(' methods ')') (() | '?' | '+' | '[' @{syntax nat}? ']') - ; - methods: (@{syntax nameref} @{syntax args} | @{syntax method}) + (',' | '|') - "} - - Proper Isar proof methods do \emph{not} admit arbitrary goal - addressing, but refer either to the first sub-goal or all sub-goals - uniformly. The goal restriction operator ``@{text "[n]"}'' - evaluates a method expression within a sandbox consisting of the - first @{text n} sub-goals (which need to exist). For example, the - method ``@{text "simp_all[3]"}'' simplifies the first three - sub-goals, while ``@{text "(rule foo, simp_all)[]"}'' simplifies all - new goals that emerge from applying rule @{text "foo"} to the - originally first one. - - Improper methods, notably tactic emulations, offer a separate - low-level goal addressing scheme as explicit argument to the - individual tactic being involved. Here ``@{text "[!]"}'' refers to - all goals, and ``@{text "[n-]"}'' to all goals starting from @{text - "n"}. - - @{rail " - @{syntax_def goal_spec}: - '[' (@{syntax nat} '-' @{syntax nat} | @{syntax nat} '-' | @{syntax nat} | '!' ) ']' - "} -*} - - -subsection {* Initial and terminal proof steps \label{sec:proof-steps} *} - -text {* - \begin{matharray}{rcl} - @{command_def "proof"} & : & @{text "proof(prove) \ proof(state)"} \\ - @{command_def "qed"} & : & @{text "proof(state) \ proof(state) | local_theory | theory"} \\ - @{command_def "by"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ - @{command_def ".."} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ - @{command_def "."} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ - @{command_def "sorry"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ - \end{matharray} - - Arbitrary goal refinement via tactics is considered harmful. - Structured proof composition in Isar admits proof methods to be - invoked in two places only. - - \begin{enumerate} - - \item An \emph{initial} refinement step @{command_ref - "proof"}~@{text "m\<^sub>1"} reduces a newly stated goal to a number - of sub-goals that are to be solved later. Facts are passed to - @{text "m\<^sub>1"} for forward chaining, if so indicated by @{text - "proof(chain)"} mode. - - \item A \emph{terminal} conclusion step @{command_ref "qed"}~@{text - "m\<^sub>2"} is intended to solve remaining goals. No facts are - passed to @{text "m\<^sub>2"}. - - \end{enumerate} - - The only other (proper) way to affect pending goals in a proof body - is by @{command_ref "show"}, which involves an explicit statement of - what is to be solved eventually. Thus we avoid the fundamental - problem of unstructured tactic scripts that consist of numerous - consecutive goal transformations, with invisible effects. - - \medskip As a general rule of thumb for good proof style, initial - proof methods should either solve the goal completely, or constitute - some well-understood reduction to new sub-goals. Arbitrary - automatic proof tools that are prone leave a large number of badly - structured sub-goals are no help in continuing the proof document in - an intelligible manner. - - Unless given explicitly by the user, the default initial method is - @{method_ref (Pure) rule} (or its classical variant @{method_ref - rule}), which applies a single standard elimination or introduction - rule according to the topmost symbol involved. There is no separate - default terminal method. Any remaining goals are always solved by - assumption in the very last step. - - @{rail " - @@{command proof} method? - ; - @@{command qed} method? - ; - @@{command \"by\"} method method? - ; - (@@{command \".\"} | @@{command \"..\"} | @@{command sorry}) - "} - - \begin{description} - - \item @{command "proof"}~@{text "m\<^sub>1"} refines the goal by proof - method @{text "m\<^sub>1"}; facts for forward chaining are passed if so - indicated by @{text "proof(chain)"} mode. - - \item @{command "qed"}~@{text "m\<^sub>2"} refines any remaining goals by - proof method @{text "m\<^sub>2"} and concludes the sub-proof by assumption. - If the goal had been @{text "show"} (or @{text "thus"}), some - pending sub-goal is solved as well by the rule resulting from the - result \emph{exported} into the enclosing goal context. Thus @{text - "qed"} may fail for two reasons: either @{text "m\<^sub>2"} fails, or the - resulting rule does not fit to any pending goal\footnote{This - includes any additional ``strong'' assumptions as introduced by - @{command "assume"}.} of the enclosing context. Debugging such a - situation might involve temporarily changing @{command "show"} into - @{command "have"}, or weakening the local context by replacing - occurrences of @{command "assume"} by @{command "presume"}. - - \item @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} is a \emph{terminal - proof}\index{proof!terminal}; it abbreviates @{command - "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"}, but with - backtracking across both methods. Debugging an unsuccessful - @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} command can be done by expanding its - definition; in many cases @{command "proof"}~@{text "m\<^sub>1"} (or even - @{text "apply"}~@{text "m\<^sub>1"}) is already sufficient to see the - problem. - - \item ``@{command ".."}'' is a \emph{default - proof}\index{proof!default}; it abbreviates @{command "by"}~@{text - "rule"}. - - \item ``@{command "."}'' is a \emph{trivial - proof}\index{proof!trivial}; it abbreviates @{command "by"}~@{text - "this"}. - - \item @{command "sorry"} is a \emph{fake proof}\index{proof!fake} - pretending to solve the pending claim without further ado. This - only works in interactive development, or if the @{ML - quick_and_dirty} flag is enabled (in ML). Facts emerging from fake - proofs are not the real thing. Internally, each theorem container - is tainted by an oracle invocation, which is indicated as ``@{text - "[!]"}'' in the printed result. - - The most important application of @{command "sorry"} is to support - experimentation and top-down proof development. - - \end{description} -*} - - -subsection {* Fundamental methods and attributes \label{sec:pure-meth-att} *} - -text {* - The following proof methods and attributes refer to basic logical - operations of Isar. Further methods and attributes are provided by - several generic and object-logic specific tools and packages (see - \chref{ch:gen-tools} and \chref{ch:hol}). - - \begin{matharray}{rcl} - @{method_def "-"} & : & @{text method} \\ - @{method_def "fact"} & : & @{text method} \\ - @{method_def "assumption"} & : & @{text method} \\ - @{method_def "this"} & : & @{text method} \\ - @{method_def (Pure) "rule"} & : & @{text method} \\ - @{attribute_def (Pure) "intro"} & : & @{text attribute} \\ - @{attribute_def (Pure) "elim"} & : & @{text attribute} \\ - @{attribute_def (Pure) "dest"} & : & @{text attribute} \\ - @{attribute_def (Pure) "rule"} & : & @{text attribute} \\[0.5ex] - @{attribute_def "OF"} & : & @{text attribute} \\ - @{attribute_def "of"} & : & @{text attribute} \\ - @{attribute_def "where"} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{method fact} @{syntax thmrefs}? - ; - @@{method (Pure) rule} @{syntax thmrefs}? - ; - rulemod: ('intro' | 'elim' | 'dest') - ((('!' | () | '?') @{syntax nat}?) | 'del') ':' @{syntax thmrefs} - ; - (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) - ('!' | () | '?') @{syntax nat}? - ; - @@{attribute (Pure) rule} 'del' - ; - @@{attribute OF} @{syntax thmrefs} - ; - @@{attribute of} @{syntax insts} ('concl' ':' @{syntax insts})? - ; - @@{attribute \"where\"} - ((@{syntax name} | @{syntax var} | @{syntax typefree} | @{syntax typevar}) '=' - (@{syntax type} | @{syntax term}) * @'and') - "} - - \begin{description} - - \item ``@{method "-"}'' (minus) does nothing but insert the forward - chaining facts as premises into the goal. Note that command - @{command_ref "proof"} without any method actually performs a single - reduction step using the @{method_ref (Pure) rule} method; thus a plain - \emph{do-nothing} proof step would be ``@{command "proof"}~@{text - "-"}'' rather than @{command "proof"} alone. - - \item @{method "fact"}~@{text "a\<^sub>1 \ a\<^sub>n"} composes some fact from - @{text "a\<^sub>1, \, a\<^sub>n"} (or implicitly from the current proof context) - modulo unification of schematic type and term variables. The rule - structure is not taken into account, i.e.\ meta-level implication is - considered atomic. This is the same principle underlying literal - facts (cf.\ \secref{sec:syn-att}): ``@{command "have"}~@{text - "\"}~@{command "by"}~@{text fact}'' is equivalent to ``@{command - "note"}~@{verbatim "`"}@{text \}@{verbatim "`"}'' provided that - @{text "\ \"} is an instance of some known @{text "\ \"} in the - proof context. - - \item @{method assumption} solves some goal by a single assumption - step. All given facts are guaranteed to participate in the - refinement; this means there may be only 0 or 1 in the first place. - Recall that @{command "qed"} (\secref{sec:proof-steps}) already - concludes any remaining sub-goals by assumption, so structured - proofs usually need not quote the @{method assumption} method at - all. - - \item @{method this} applies all of the current facts directly as - rules. Recall that ``@{command "."}'' (dot) abbreviates ``@{command - "by"}~@{text this}''. - - \item @{method (Pure) rule}~@{text "a\<^sub>1 \ a\<^sub>n"} applies some rule given as - argument in backward manner; facts are used to reduce the rule - before applying it to the goal. Thus @{method (Pure) rule} without facts - is plain introduction, while with facts it becomes elimination. - - When no arguments are given, the @{method (Pure) rule} method tries to pick - appropriate rules automatically, as declared in the current context - using the @{attribute (Pure) intro}, @{attribute (Pure) elim}, - @{attribute (Pure) dest} attributes (see below). This is the - default behavior of @{command "proof"} and ``@{command ".."}'' - (double-dot) steps (see \secref{sec:proof-steps}). - - \item @{attribute (Pure) intro}, @{attribute (Pure) elim}, and - @{attribute (Pure) dest} declare introduction, elimination, and - destruct rules, to be used with method @{method (Pure) rule}, and similar - tools. Note that the latter will ignore rules declared with - ``@{text "?"}'', while ``@{text "!"}'' are used most aggressively. - - The classical reasoner (see \secref{sec:classical}) introduces its - own variants of these attributes; use qualified names to access the - present versions of Isabelle/Pure, i.e.\ @{attribute (Pure) - "Pure.intro"}. - - \item @{attribute (Pure) rule}~@{text del} undeclares introduction, - elimination, or destruct rules. - - \item @{attribute OF}~@{text "a\<^sub>1 \ a\<^sub>n"} applies some theorem to all - of the given rules @{text "a\<^sub>1, \, a\<^sub>n"} in canonical right-to-left - order, which means that premises stemming from the @{text "a\<^sub>i"} - emerge in parallel in the result, without interfering with each - other. In many practical situations, the @{text "a\<^sub>i"} do not have - premises themselves, so @{text "rule [OF a\<^sub>1 \ a\<^sub>n]"} can be actually - read as functional application (modulo unification). - - Argument positions may be effectively skipped by using ``@{text _}'' - (underscore), which refers to the propositional identity rule in the - Pure theory. - - \item @{attribute of}~@{text "t\<^sub>1 \ t\<^sub>n"} performs positional - instantiation of term variables. The terms @{text "t\<^sub>1, \, t\<^sub>n"} are - substituted for any schematic variables occurring in a theorem from - left to right; ``@{text _}'' (underscore) indicates to skip a - position. Arguments following a ``@{text "concl:"}'' specification - refer to positions of the conclusion of a rule. - - \item @{attribute "where"}~@{text "x\<^sub>1 = t\<^sub>1 \ \ x\<^sub>n = t\<^sub>n"} - performs named instantiation of schematic type and term variables - occurring in a theorem. Schematic variables have to be specified on - the left-hand side (e.g.\ @{text "?x1.3"}). The question mark may - be omitted if the variable name is a plain identifier without index. - As type instantiations are inferred from term instantiations, - explicit type instantiations are seldom necessary. - - \end{description} -*} - - -subsection {* Emulating tactic scripts \label{sec:tactic-commands} *} - -text {* - The Isar provides separate commands to accommodate tactic-style - proof scripts within the same system. While being outside the - orthodox Isar proof language, these might come in handy for - interactive exploration and debugging, or even actual tactical proof - within new-style theories (to benefit from document preparation, for - example). See also \secref{sec:tactics} for actual tactics, that - have been encapsulated as proof methods. Proper proof methods may - be used in scripts, too. - - \begin{matharray}{rcl} - @{command_def "apply"}@{text "\<^sup>*"} & : & @{text "proof(prove) \ proof(prove)"} \\ - @{command_def "apply_end"}@{text "\<^sup>*"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "done"}@{text "\<^sup>*"} & : & @{text "proof(prove) \ proof(state) | local_theory | theory"} \\ - @{command_def "defer"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ - @{command_def "prefer"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ - @{command_def "back"}@{text "\<^sup>*"} & : & @{text "proof \ proof"} \\ - \end{matharray} - - @{rail " - ( @@{command apply} | @@{command apply_end} ) @{syntax method} - ; - @@{command defer} @{syntax nat}? - ; - @@{command prefer} @{syntax nat} - "} - - \begin{description} - - \item @{command "apply"}~@{text m} applies proof method @{text m} in - initial position, but unlike @{command "proof"} it retains ``@{text - "proof(prove)"}'' mode. Thus consecutive method applications may be - given just as in tactic scripts. - - Facts are passed to @{text m} as indicated by the goal's - forward-chain mode, and are \emph{consumed} afterwards. Thus any - further @{command "apply"} command would always work in a purely - backward manner. - - \item @{command "apply_end"}~@{text "m"} applies proof method @{text - m} as if in terminal position. Basically, this simulates a - multi-step tactic script for @{command "qed"}, but may be given - anywhere within the proof body. - - No facts are passed to @{text m} here. Furthermore, the static - context is that of the enclosing goal (as for actual @{command - "qed"}). Thus the proof method may not refer to any assumptions - introduced in the current body, for example. - - \item @{command "done"} completes a proof script, provided that the - current goal state is solved completely. Note that actual - structured proof commands (e.g.\ ``@{command "."}'' or @{command - "sorry"}) may be used to conclude proof scripts as well. - - \item @{command "defer"}~@{text n} and @{command "prefer"}~@{text n} - shuffle the list of pending goals: @{command "defer"} puts off - sub-goal @{text n} to the end of the list (@{text "n = 1"} by - default), while @{command "prefer"} brings sub-goal @{text n} to the - front. - - \item @{command "back"} does back-tracking over the result sequence - of the latest proof command. Basically, any proof command may - return multiple results. - - \end{description} - - Any proper Isar proof method may be used with tactic script commands - such as @{command "apply"}. A few additional emulations of actual - tactics are provided as well; these would be never used in actual - structured proofs, of course. -*} - - -subsection {* Defining proof methods *} - -text {* - \begin{matharray}{rcl} - @{command_def "method_setup"} & : & @{text "theory \ theory"} \\ - \end{matharray} - - @{rail " - @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}? - ; - "} - - \begin{description} - - \item @{command "method_setup"}~@{text "name = text description"} - defines a proof method in the current theory. The given @{text - "text"} has to be an ML expression of type - @{ML_type "(Proof.context -> Proof.method) context_parser"}, cf.\ - basic parsers defined in structure @{ML_struct Args} and @{ML_struct - Attrib}. There are also combinators like @{ML METHOD} and @{ML - SIMPLE_METHOD} to turn certain tactic forms into official proof - methods; the primed versions refer to tactics with explicit goal - addressing. - - Here are some example method definitions: - - \end{description} -*} - - method_setup my_method1 = {* - Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac))) - *} "my first method (without any arguments)" - - method_setup my_method2 = {* - Scan.succeed (fn ctxt: Proof.context => - SIMPLE_METHOD' (fn i: int => no_tac)) - *} "my second method (with context)" - - method_setup my_method3 = {* - Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context => - SIMPLE_METHOD' (fn i: int => no_tac)) - *} "my third method (with theorem arguments and context)" - - -section {* Generalized elimination \label{sec:obtain} *} - -text {* - \begin{matharray}{rcl} - @{command_def "obtain"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ - @{command_def "guess"}@{text "\<^sup>*"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ - \end{matharray} - - Generalized elimination means that additional elements with certain - properties may be introduced in the current context, by virtue of a - locally proven ``soundness statement''. Technically speaking, the - @{command "obtain"} language element is like a declaration of - @{command "fix"} and @{command "assume"} (see also see - \secref{sec:proof-context}), together with a soundness proof of its - additional claim. According to the nature of existential reasoning, - assumptions get eliminated from any result exported from the context - later, provided that the corresponding parameters do \emph{not} - occur in the conclusion. - - @{rail " - @@{command obtain} @{syntax parname}? (@{syntax vars} + @'and') - @'where' (@{syntax props} + @'and') - ; - @@{command guess} (@{syntax vars} + @'and') - "} - - The derived Isar command @{command "obtain"} is defined as follows - (where @{text "b\<^sub>1, \, b\<^sub>k"} shall refer to (optional) - facts indicated for forward chaining). - \begin{matharray}{l} - @{text "\using b\<^sub>1 \ b\<^sub>k\"}~~@{command "obtain"}~@{text "x\<^sub>1 \ x\<^sub>m \ a: \\<^sub>1 \ \\<^sub>n \proof\ \"} \\[1ex] - \quad @{command "have"}~@{text "\thesis. (\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis) \ thesis"} \\ - \quad @{command "proof"}~@{method succeed} \\ - \qquad @{command "fix"}~@{text thesis} \\ - \qquad @{command "assume"}~@{text "that [Pure.intro?]: \x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ thesis"} \\ - \qquad @{command "then"}~@{command "show"}~@{text thesis} \\ - \quad\qquad @{command "apply"}~@{text -} \\ - \quad\qquad @{command "using"}~@{text "b\<^sub>1 \ b\<^sub>k \proof\"} \\ - \quad @{command "qed"} \\ - \quad @{command "fix"}~@{text "x\<^sub>1 \ x\<^sub>m"}~@{command "assume"}@{text "\<^sup>* a: \\<^sub>1 \ \\<^sub>n"} \\ - \end{matharray} - - Typically, the soundness proof is relatively straight-forward, often - just by canonical automated tools such as ``@{command "by"}~@{text - simp}'' or ``@{command "by"}~@{text blast}''. Accordingly, the - ``@{text that}'' reduction above is declared as simplification and - introduction rule. - - In a sense, @{command "obtain"} represents at the level of Isar - proofs what would be meta-logical existential quantifiers and - conjunctions. This concept has a broad range of useful - applications, ranging from plain elimination (or introduction) of - object-level existential and conjunctions, to elimination over - results of symbolic evaluation of recursive definitions, for - example. Also note that @{command "obtain"} without parameters acts - much like @{command "have"}, where the result is treated as a - genuine assumption. - - An alternative name to be used instead of ``@{text that}'' above may - be given in parentheses. - - \medskip The improper variant @{command "guess"} is similar to - @{command "obtain"}, but derives the obtained statement from the - course of reasoning! The proof starts with a fixed goal @{text - thesis}. The subsequent proof may refine this to anything of the - form like @{text "\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ - \\<^sub>n \ thesis"}, but must not introduce new subgoals. The - final goal state is then used as reduction rule for the obtain - scheme described above. Obtained parameters @{text "x\<^sub>1, \, - x\<^sub>m"} are marked as internal by default, which prevents the - proof context from being polluted by ad-hoc variables. The variable - names and type constraints given as arguments for @{command "guess"} - specify a prefix of obtained parameters explicitly in the text. - - It is important to note that the facts introduced by @{command - "obtain"} and @{command "guess"} may not be polymorphic: any - type-variables occurring here are fixed in the present context! -*} - - -section {* Calculational reasoning \label{sec:calculation} *} - -text {* - \begin{matharray}{rcl} - @{command_def "also"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "finally"} & : & @{text "proof(state) \ proof(chain)"} \\ - @{command_def "moreover"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "ultimately"} & : & @{text "proof(state) \ proof(chain)"} \\ - @{command_def "print_trans_rules"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute trans} & : & @{text attribute} \\ - @{attribute sym} & : & @{text attribute} \\ - @{attribute symmetric} & : & @{text attribute} \\ - \end{matharray} - - Calculational proof is forward reasoning with implicit application - of transitivity rules (such those of @{text "="}, @{text "\"}, - @{text "<"}). Isabelle/Isar maintains an auxiliary fact register - @{fact_ref calculation} for accumulating results obtained by - transitivity composed with the current result. Command @{command - "also"} updates @{fact calculation} involving @{fact this}, while - @{command "finally"} exhibits the final @{fact calculation} by - forward chaining towards the next goal statement. Both commands - require valid current facts, i.e.\ may occur only after commands - that produce theorems such as @{command "assume"}, @{command - "note"}, or some finished proof of @{command "have"}, @{command - "show"} etc. The @{command "moreover"} and @{command "ultimately"} - commands are similar to @{command "also"} and @{command "finally"}, - but only collect further results in @{fact calculation} without - applying any rules yet. - - Also note that the implicit term abbreviation ``@{text "\"}'' has - its canonical application with calculational proofs. It refers to - the argument of the preceding statement. (The argument of a curried - infix expression happens to be its right-hand side.) - - Isabelle/Isar calculations are implicitly subject to block structure - in the sense that new threads of calculational reasoning are - commenced for any new block (as opened by a local goal, for - example). This means that, apart from being able to nest - calculations, there is no separate \emph{begin-calculation} command - required. - - \medskip The Isar calculation proof commands may be defined as - follows:\footnote{We suppress internal bookkeeping such as proper - handling of block-structure.} - - \begin{matharray}{rcl} - @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\ - @{command "also"}@{text "\<^sub>n+1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex] - @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] - @{command "moreover"} & \equiv & @{command "note"}~@{text "calculation = calculation this"} \\ - @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~@{text calculation} \\ - \end{matharray} - - @{rail " - (@@{command also} | @@{command finally}) ('(' @{syntax thmrefs} ')')? - ; - @@{attribute trans} (() | 'add' | 'del') - "} - - \begin{description} - - \item @{command "also"}~@{text "(a\<^sub>1 \ a\<^sub>n)"} maintains the auxiliary - @{fact calculation} register as follows. The first occurrence of - @{command "also"} in some calculational thread initializes @{fact - calculation} by @{fact this}. Any subsequent @{command "also"} on - the same level of block-structure updates @{fact calculation} by - some transitivity rule applied to @{fact calculation} and @{fact - this} (in that order). Transitivity rules are picked from the - current context, unless alternative rules are given as explicit - arguments. - - \item @{command "finally"}~@{text "(a\<^sub>1 \ a\<^sub>n)"} maintaining @{fact - calculation} in the same way as @{command "also"}, and concludes the - current calculational thread. The final result is exhibited as fact - for forward chaining towards the next goal. Basically, @{command - "finally"} just abbreviates @{command "also"}~@{command - "from"}~@{fact calculation}. Typical idioms for concluding - calculational proofs are ``@{command "finally"}~@{command - "show"}~@{text ?thesis}~@{command "."}'' and ``@{command - "finally"}~@{command "have"}~@{text \}~@{command "."}''. - - \item @{command "moreover"} and @{command "ultimately"} are - analogous to @{command "also"} and @{command "finally"}, but collect - results only, without applying rules. - - \item @{command "print_trans_rules"} prints the list of transitivity - rules (for calculational commands @{command "also"} and @{command - "finally"}) and symmetry rules (for the @{attribute symmetric} - operation and single step elimination patters) of the current - context. - - \item @{attribute trans} declares theorems as transitivity rules. - - \item @{attribute sym} declares symmetry rules, as well as - @{attribute "Pure.elim"}@{text "?"} rules. - - \item @{attribute symmetric} resolves a theorem with some rule - declared as @{attribute sym} in the current context. For example, - ``@{command "assume"}~@{text "[symmetric]: x = y"}'' produces a - swapped fact derived from that assumption. - - In structured proof texts it is often more appropriate to use an - explicit single-step elimination proof, such as ``@{command - "assume"}~@{text "x = y"}~@{command "then"}~@{command "have"}~@{text - "y = x"}~@{command ".."}''. - - \end{description} -*} - - -section {* Proof by cases and induction \label{sec:cases-induct} *} - -subsection {* Rule contexts *} - -text {* - \begin{matharray}{rcl} - @{command_def "case"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "print_cases"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def case_names} & : & @{text attribute} \\ - @{attribute_def case_conclusion} & : & @{text attribute} \\ - @{attribute_def params} & : & @{text attribute} \\ - @{attribute_def consumes} & : & @{text attribute} \\ - \end{matharray} - - The puristic way to build up Isar proof contexts is by explicit - language elements like @{command "fix"}, @{command "assume"}, - @{command "let"} (see \secref{sec:proof-context}). This is adequate - for plain natural deduction, but easily becomes unwieldy in concrete - verification tasks, which typically involve big induction rules with - several cases. - - The @{command "case"} command provides a shorthand to refer to a - local context symbolically: certain proof methods provide an - environment of named ``cases'' of the form @{text "c: x\<^sub>1, \, - x\<^sub>m, \\<^sub>1, \, \\<^sub>n"}; the effect of ``@{command - "case"}~@{text c}'' is then equivalent to ``@{command "fix"}~@{text - "x\<^sub>1 \ x\<^sub>m"}~@{command "assume"}~@{text "c: \\<^sub>1 \ - \\<^sub>n"}''. Term bindings may be covered as well, notably - @{variable ?case} for the main conclusion. - - By default, the ``terminology'' @{text "x\<^sub>1, \, x\<^sub>m"} of - a case value is marked as hidden, i.e.\ there is no way to refer to - such parameters in the subsequent proof text. After all, original - rule parameters stem from somewhere outside of the current proof - text. By using the explicit form ``@{command "case"}~@{text "(c - y\<^sub>1 \ y\<^sub>m)"}'' instead, the proof author is able to - chose local names that fit nicely into the current context. - - \medskip It is important to note that proper use of @{command - "case"} does not provide means to peek at the current goal state, - which is not directly observable in Isar! Nonetheless, goal - refinement commands do provide named cases @{text "goal\<^sub>i"} - for each subgoal @{text "i = 1, \, n"} of the resulting goal state. - Using this extra feature requires great care, because some bits of - the internal tactical machinery intrude the proof text. In - particular, parameter names stemming from the left-over of automated - reasoning tools are usually quite unpredictable. - - Under normal circumstances, the text of cases emerge from standard - elimination or induction rules, which in turn are derived from - previous theory specifications in a canonical way (say from - @{command "inductive"} definitions). - - \medskip Proper cases are only available if both the proof method - and the rules involved support this. By using appropriate - attributes, case names, conclusions, and parameters may be also - declared by hand. Thus variant versions of rules that have been - derived manually become ready to use in advanced case analysis - later. - - @{rail " - @@{command case} (caseref | '(' caseref (('_' | @{syntax name}) +) ')') - ; - caseref: nameref attributes? - ; - - @@{attribute case_names} ((@{syntax name} ( '[' (('_' | @{syntax name}) +) ']' ) ? ) +) - ; - @@{attribute case_conclusion} @{syntax name} (@{syntax name} * ) - ; - @@{attribute params} ((@{syntax name} * ) + @'and') - ; - @@{attribute consumes} @{syntax nat}? - "} - - \begin{description} - - \item @{command "case"}~@{text "(c x\<^sub>1 \ x\<^sub>m)"} invokes a named local - context @{text "c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>m"}, as provided by an - appropriate proof method (such as @{method_ref cases} and - @{method_ref induct}). The command ``@{command "case"}~@{text "(c - x\<^sub>1 \ x\<^sub>m)"}'' abbreviates ``@{command "fix"}~@{text "x\<^sub>1 \ - x\<^sub>m"}~@{command "assume"}~@{text "c: \\<^sub>1 \ \\<^sub>n"}''. - - \item @{command "print_cases"} prints all local contexts of the - current state, using Isar proof language notation. - - \item @{attribute case_names}~@{text "c\<^sub>1 \ c\<^sub>k"} declares names for - the local contexts of premises of a theorem; @{text "c\<^sub>1, \, c\<^sub>k"} - refers to the \emph{prefix} of the list of premises. Each of the - cases @{text "c\<^isub>i"} can be of the form @{text "c[h\<^isub>1 \ h\<^isub>n]"} where - the @{text "h\<^isub>1 \ h\<^isub>n"} are the names of the hypotheses in case @{text "c\<^isub>i"} - from left to right. - - \item @{attribute case_conclusion}~@{text "c d\<^sub>1 \ d\<^sub>k"} declares - names for the conclusions of a named premise @{text c}; here @{text - "d\<^sub>1, \, d\<^sub>k"} refers to the prefix of arguments of a logical formula - built by nesting a binary connective (e.g.\ @{text "\"}). - - Note that proof methods such as @{method induct} and @{method - coinduct} already provide a default name for the conclusion as a - whole. The need to name subformulas only arises with cases that - split into several sub-cases, as in common co-induction rules. - - \item @{attribute params}~@{text "p\<^sub>1 \ p\<^sub>m \ \ q\<^sub>1 \ q\<^sub>n"} renames - the innermost parameters of premises @{text "1, \, n"} of some - theorem. An empty list of names may be given to skip positions, - leaving the present parameters unchanged. - - Note that the default usage of case rules does \emph{not} directly - expose parameters to the proof context. - - \item @{attribute consumes}~@{text n} declares the number of ``major - premises'' of a rule, i.e.\ the number of facts to be consumed when - it is applied by an appropriate proof method. The default value of - @{attribute consumes} is @{text "n = 1"}, which is appropriate for - the usual kind of cases and induction rules for inductive sets (cf.\ - \secref{sec:hol-inductive}). Rules without any @{attribute - consumes} declaration given are treated as if @{attribute - consumes}~@{text 0} had been specified. - - Note that explicit @{attribute consumes} declarations are only - rarely needed; this is already taken care of automatically by the - higher-level @{attribute cases}, @{attribute induct}, and - @{attribute coinduct} declarations. - - \end{description} -*} - - -subsection {* Proof methods *} - -text {* - \begin{matharray}{rcl} - @{method_def cases} & : & @{text method} \\ - @{method_def induct} & : & @{text method} \\ - @{method_def induction} & : & @{text method} \\ - @{method_def coinduct} & : & @{text method} \\ - \end{matharray} - - The @{method cases}, @{method induct}, @{method induction}, - and @{method coinduct} - methods provide a uniform interface to common proof techniques over - datatypes, inductive predicates (or sets), recursive functions etc. - The corresponding rules may be specified and instantiated in a - casual manner. Furthermore, these methods provide named local - contexts that may be invoked via the @{command "case"} proof command - within the subsequent proof text. This accommodates compact proof - texts even when reasoning about large specifications. - - The @{method induct} method also provides some additional - infrastructure in order to be applicable to structure statements - (either using explicit meta-level connectives, or including facts - and parameters separately). This avoids cumbersome encoding of - ``strengthened'' inductive statements within the object-logic. - - Method @{method induction} differs from @{method induct} only in - the names of the facts in the local context invoked by the @{command "case"} - command. - - @{rail " - @@{method cases} ('(' 'no_simp' ')')? \\ - (@{syntax insts} * @'and') rule? - ; - (@@{method induct} | @@{method induction}) ('(' 'no_simp' ')')? (definsts * @'and') \\ arbitrary? taking? rule? - ; - @@{method coinduct} @{syntax insts} taking rule? - ; - - rule: ('type' | 'pred' | 'set') ':' (@{syntax nameref} +) | 'rule' ':' (@{syntax thmref} +) - ; - definst: @{syntax name} ('==' | '\') @{syntax term} | '(' @{syntax term} ')' | @{syntax inst} - ; - definsts: ( definst * ) - ; - arbitrary: 'arbitrary' ':' ((@{syntax term} * ) @'and' +) - ; - taking: 'taking' ':' @{syntax insts} - "} - - \begin{description} - - \item @{method cases}~@{text "insts R"} applies method @{method - rule} with an appropriate case distinction theorem, instantiated to - the subjects @{text insts}. Symbolic case names are bound according - to the rule's local contexts. - - The rule is determined as follows, according to the facts and - arguments passed to the @{method cases} method: - - \medskip - \begin{tabular}{llll} - facts & & arguments & rule \\\hline - & @{method cases} & & classical case split \\ - & @{method cases} & @{text t} & datatype exhaustion (type of @{text t}) \\ - @{text "\ A t"} & @{method cases} & @{text "\"} & inductive predicate/set elimination (of @{text A}) \\ - @{text "\"} & @{method cases} & @{text "\ rule: R"} & explicit rule @{text R} \\ - \end{tabular} - \medskip - - Several instantiations may be given, referring to the \emph{suffix} - of premises of the case rule; within each premise, the \emph{prefix} - of variables is instantiated. In most situations, only a single - term needs to be specified; this refers to the first variable of the - last premise (it is usually the same for all cases). The @{text - "(no_simp)"} option can be used to disable pre-simplification of - cases (see the description of @{method induct} below for details). - - \item @{method induct}~@{text "insts R"} and - @{method induction}~@{text "insts R"} are analogous to the - @{method cases} method, but refer to induction rules, which are - determined as follows: - - \medskip - \begin{tabular}{llll} - facts & & arguments & rule \\\hline - & @{method induct} & @{text "P x"} & datatype induction (type of @{text x}) \\ - @{text "\ A x"} & @{method induct} & @{text "\"} & predicate/set induction (of @{text A}) \\ - @{text "\"} & @{method induct} & @{text "\ rule: R"} & explicit rule @{text R} \\ - \end{tabular} - \medskip - - Several instantiations may be given, each referring to some part of - a mutual inductive definition or datatype --- only related partial - induction rules may be used together, though. Any of the lists of - terms @{text "P, x, \"} refers to the \emph{suffix} of variables - present in the induction rule. This enables the writer to specify - only induction variables, or both predicates and variables, for - example. - - Instantiations may be definitional: equations @{text "x \ t"} - introduce local definitions, which are inserted into the claim and - discharged after applying the induction rule. Equalities reappear - in the inductive cases, but have been transformed according to the - induction principle being involved here. In order to achieve - practically useful induction hypotheses, some variables occurring in - @{text t} need to be fixed (see below). Instantiations of the form - @{text t}, where @{text t} is not a variable, are taken as a - shorthand for \mbox{@{text "x \ t"}}, where @{text x} is a fresh - variable. If this is not intended, @{text t} has to be enclosed in - parentheses. By default, the equalities generated by definitional - instantiations are pre-simplified using a specific set of rules, - usually consisting of distinctness and injectivity theorems for - datatypes. This pre-simplification may cause some of the parameters - of an inductive case to disappear, or may even completely delete - some of the inductive cases, if one of the equalities occurring in - their premises can be simplified to @{text False}. The @{text - "(no_simp)"} option can be used to disable pre-simplification. - Additional rules to be used in pre-simplification can be declared - using the @{attribute_def induct_simp} attribute. - - The optional ``@{text "arbitrary: x\<^sub>1 \ x\<^sub>m"}'' - specification generalizes variables @{text "x\<^sub>1, \, - x\<^sub>m"} of the original goal before applying induction. One can - separate variables by ``@{text "and"}'' to generalize them in other - goals then the first. Thus induction hypotheses may become - sufficiently general to get the proof through. Together with - definitional instantiations, one may effectively perform induction - over expressions of a certain structure. - - The optional ``@{text "taking: t\<^sub>1 \ t\<^sub>n"}'' - specification provides additional instantiations of a prefix of - pending variables in the rule. Such schematic induction rules - rarely occur in practice, though. - - \item @{method coinduct}~@{text "inst R"} is analogous to the - @{method induct} method, but refers to coinduction rules, which are - determined as follows: - - \medskip - \begin{tabular}{llll} - goal & & arguments & rule \\\hline - & @{method coinduct} & @{text x} & type coinduction (type of @{text x}) \\ - @{text "A x"} & @{method coinduct} & @{text "\"} & predicate/set coinduction (of @{text A}) \\ - @{text "\"} & @{method coinduct} & @{text "\ rule: R"} & explicit rule @{text R} \\ - \end{tabular} - - Coinduction is the dual of induction. Induction essentially - eliminates @{text "A x"} towards a generic result @{text "P x"}, - while coinduction introduces @{text "A x"} starting with @{text "B - x"}, for a suitable ``bisimulation'' @{text B}. The cases of a - coinduct rule are typically named after the predicates or sets being - covered, while the conclusions consist of several alternatives being - named after the individual destructor patterns. - - The given instantiation refers to the \emph{suffix} of variables - occurring in the rule's major premise, or conclusion if unavailable. - An additional ``@{text "taking: t\<^sub>1 \ t\<^sub>n"}'' - specification may be required in order to specify the bisimulation - to be used in the coinduction step. - - \end{description} - - Above methods produce named local contexts, as determined by the - instantiated rule as given in the text. Beyond that, the @{method - induct} and @{method coinduct} methods guess further instantiations - from the goal specification itself. Any persisting unresolved - schematic variables of the resulting rule will render the the - corresponding case invalid. The term binding @{variable ?case} for - the conclusion will be provided with each case, provided that term - is fully specified. - - The @{command "print_cases"} command prints all named cases present - in the current proof state. - - \medskip Despite the additional infrastructure, both @{method cases} - and @{method coinduct} merely apply a certain rule, after - instantiation, while conforming due to the usual way of monotonic - natural deduction: the context of a structured statement @{text - "\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ \"} - reappears unchanged after the case split. - - The @{method induct} method is fundamentally different in this - respect: the meta-level structure is passed through the - ``recursive'' course involved in the induction. Thus the original - statement is basically replaced by separate copies, corresponding to - the induction hypotheses and conclusion; the original goal context - is no longer available. Thus local assumptions, fixed parameters - and definitions effectively participate in the inductive rephrasing - of the original statement. - - In @{method induct} proofs, local assumptions introduced by cases are split - into two different kinds: @{text hyps} stemming from the rule and - @{text prems} from the goal statement. This is reflected in the - extracted cases accordingly, so invoking ``@{command "case"}~@{text - c}'' will provide separate facts @{text c.hyps} and @{text c.prems}, - as well as fact @{text c} to hold the all-inclusive list. - - In @{method induction} proofs, local assumptions introduced by cases are - split into three different kinds: @{text IH}, the induction hypotheses, - @{text hyps}, the remaining hypotheses stemming from the rule, and - @{text prems}, the assumptions from the goal statement. The names are - @{text c.IH}, @{text c.hyps} and @{text c.prems}, as above. - - - \medskip Facts presented to either method are consumed according to - the number of ``major premises'' of the rule involved, which is - usually 0 for plain cases and induction rules of datatypes etc.\ and - 1 for rules of inductive predicates or sets and the like. The - remaining facts are inserted into the goal verbatim before the - actual @{text cases}, @{text induct}, or @{text coinduct} rule is - applied. -*} - - -subsection {* Declaring rules *} - -text {* - \begin{matharray}{rcl} - @{command_def "print_induct_rules"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{attribute_def cases} & : & @{text attribute} \\ - @{attribute_def induct} & : & @{text attribute} \\ - @{attribute_def coinduct} & : & @{text attribute} \\ - \end{matharray} - - @{rail " - @@{attribute cases} spec - ; - @@{attribute induct} spec - ; - @@{attribute coinduct} spec - ; - - spec: (('type' | 'pred' | 'set') ':' @{syntax nameref}) | 'del' - "} - - \begin{description} - - \item @{command "print_induct_rules"} prints cases and induct rules - for predicates (or sets) and types of the current context. - - \item @{attribute cases}, @{attribute induct}, and @{attribute - coinduct} (as attributes) declare rules for reasoning about - (co)inductive predicates (or sets) and types, using the - corresponding methods of the same name. Certain definitional - packages of object-logics usually declare emerging cases and - induction rules as expected, so users rarely need to intervene. - - Rules may be deleted via the @{text "del"} specification, which - covers all of the @{text "type"}/@{text "pred"}/@{text "set"} - sub-categories simultaneously. For example, @{attribute - cases}~@{text del} removes any @{attribute cases} rules declared for - some type, predicate, or set. - - Manual rule declarations usually refer to the @{attribute - case_names} and @{attribute params} attributes to adjust names of - cases and parameters of a rule; the @{attribute consumes} - declaration is taken care of automatically: @{attribute - consumes}~@{text 0} is specified for ``type'' rules and @{attribute - consumes}~@{text 1} for ``predicate'' / ``set'' rules. - - \end{description} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Quick_Reference.thy --- a/doc-src/IsarRef/Thy/Quick_Reference.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,231 +0,0 @@ -theory Quick_Reference -imports Base Main -begin - -chapter {* Isabelle/Isar quick reference \label{ap:refcard} *} - -section {* Proof commands *} - -subsection {* Primitives and basic syntax *} - -text {* - \begin{tabular}{ll} - @{command "fix"}~@{text x} & augment context by @{text "\x. \"} \\ - @{command "assume"}~@{text "a: \"} & augment context by @{text "\ \ \"} \\ - @{command "then"} & indicate forward chaining of facts \\ - @{command "have"}~@{text "a: \"} & prove local result \\ - @{command "show"}~@{text "a: \"} & prove local result, refining some goal \\ - @{command "using"}~@{text a} & indicate use of additional facts \\ - @{command "unfolding"}~@{text a} & unfold definitional equations \\ - @{command "proof"}~@{text "m\<^sub>1"}~\dots~@{command "qed"}~@{text "m\<^sub>2"} & indicate proof structure and refinements \\ - @{command "{"}~@{text "\"}~@{command "}"} & indicate explicit blocks \\ - @{command "next"} & switch blocks \\ - @{command "note"}~@{text "a = b"} & reconsider facts \\ - @{command "let"}~@{text "p = t"} & abbreviate terms by higher-order matching \\ - @{command "write"}~@{text "c (mx)"} & declare local mixfix syntax \\ - \end{tabular} - - \medskip - - \begin{tabular}{rcl} - @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\ - & @{text "|"} & @{text "prfx\<^sup>*"}~@{command "done"} \\ - @{text prfx} & = & @{command "apply"}~@{text method} \\ - & @{text "|"} & @{command "using"}~@{text "facts"} \\ - & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ - @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ - & @{text "|"} & @{command "next"} \\ - & @{text "|"} & @{command "note"}~@{text "name = facts"} \\ - & @{text "|"} & @{command "let"}~@{text "term = term"} \\ - & @{text "|"} & @{command "write"}~@{text "name (mixfix)"} \\ - & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ - & @{text "|"} & @{command "assume"}~@{text "name: props"} \\ - & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ - @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ - & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ - \end{tabular} -*} - - -subsection {* Abbreviations and synonyms *} - -text {* - \begin{tabular}{rcl} - @{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} & @{text "\"} & - @{command "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"} \\ - @{command ".."} & @{text "\"} & @{command "by"}~@{text rule} \\ - @{command "."} & @{text "\"} & @{command "by"}~@{text this} \\ - @{command "hence"} & @{text "\"} & @{command "then"}~@{command "have"} \\ - @{command "thus"} & @{text "\"} & @{command "then"}~@{command "show"} \\ - @{command "from"}~@{text a} & @{text "\"} & @{command "note"}~@{text a}~@{command "then"} \\ - @{command "with"}~@{text a} & @{text "\"} & @{command "from"}~@{text "a \ this"} \\ - @{command "from"}~@{text this} & @{text "\"} & @{command "then"} \\ - @{command "from"}~@{text this}~@{command "have"} & @{text "\"} & @{command "hence"} \\ - @{command "from"}~@{text this}~@{command "show"} & @{text "\"} & @{command "thus"} \\ - \end{tabular} -*} - - -subsection {* Derived elements *} - -text {* - \begin{tabular}{rcl} - @{command "also"}@{text "\<^sub>0"} & @{text "\"} & - @{command "note"}~@{text "calculation = this"} \\ - @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & @{text "\"} & - @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\ - @{command "finally"} & @{text "\"} & - @{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] - @{command "moreover"} & @{text "\"} & - @{command "note"}~@{text "calculation = calculation this"} \\ - @{command "ultimately"} & @{text "\"} & - @{command "moreover"}~@{command "from"}~@{text calculation} \\[0.5ex] - @{command "presume"}~@{text "a: \"} & @{text "\"} & - @{command "assume"}~@{text "a: \"} \\ - @{command "def"}~@{text "a: x \ t"} & @{text "\"} & - @{command "fix"}~@{text x}~@{command "assume"}~@{text "a: x \ t"} \\ - @{command "obtain"}~@{text "x \ a: \"} & @{text "\"} & - @{text "\"}~@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: \"} \\ - @{command "case"}~@{text c} & @{text "\"} & - @{command "fix"}~@{text x}~@{command "assume"}~@{text "c: \"} \\ - @{command "sorry"} & @{text "\"} & - @{command "by"}~@{text cheating} \\ - \end{tabular} -*} - - -subsection {* Diagnostic commands *} - -text {* - \begin{tabular}{ll} - @{command "pr"} & print current state \\ - @{command "thm"}~@{text a} & print fact \\ - @{command "prop"}~@{text \} & print proposition \\ - @{command "term"}~@{text t} & print term \\ - @{command "typ"}~@{text \} & print type \\ - \end{tabular} -*} - - -section {* Proof methods *} - -text {* - \begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] - @{method assumption} & apply some assumption \\ - @{method this} & apply current facts \\ - @{method rule}~@{text a} & apply some rule \\ - @{method rule} & apply standard rule (default for @{command "proof"}) \\ - @{method contradiction} & apply @{text "\"} elimination rule (any order) \\ - @{method cases}~@{text t} & case analysis (provides cases) \\ - @{method induct}~@{text x} & proof by induction (provides cases) \\[2ex] - - \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] - @{method "-"} & no rules \\ - @{method intro}~@{text a} & introduction rules \\ - @{method intro_classes} & class introduction rules \\ - @{method elim}~@{text a} & elimination rules \\ - @{method unfold}~@{text a} & definitional rewrite rules \\[2ex] - - \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] - @{method iprover} & intuitionistic proof search \\ - @{method blast}, @{method fast} & Classical Reasoner \\ - @{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\ - @{method auto}, @{method force} & Simplifier + Classical Reasoner \\ - @{method arith} & Arithmetic procedures \\ - \end{tabular} -*} - - -section {* Attributes *} - -text {* - \begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Rules}} \\[0.5ex] - @{attribute OF}~@{text a} & rule resolved with facts (skipping ``@{text _}'') \\ - @{attribute of}~@{text t} & rule instantiated with terms (skipping ``@{text _}'') \\ - @{attribute "where"}~@{text "x = t"} & rule instantiated with terms, by variable name \\ - @{attribute symmetric} & resolution with symmetry rule \\ - @{attribute THEN}~@{text b} & resolution with another rule \\ - @{attribute rule_format} & result put into standard rule format \\ - @{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex] - - \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] - @{attribute simp} & Simplifier rule \\ - @{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\ - @{attribute iff} & Simplifier + Classical Reasoner rule \\ - @{attribute split} & case split rule \\ - @{attribute trans} & transitivity rule \\ - @{attribute sym} & symmetry rule \\ - \end{tabular} -*} - - -section {* Rule declarations and methods *} - -text {* - \begin{tabular}{l|lllll} - & @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\ - & & & @{method fast} & @{method simp_all} & @{method force} \\ - \hline - @{attribute Pure.elim}@{text "!"} @{attribute Pure.intro}@{text "!"} - & @{text "\"} & @{text "\"} \\ - @{attribute Pure.elim} @{attribute Pure.intro} - & @{text "\"} & @{text "\"} \\ - @{attribute elim}@{text "!"} @{attribute intro}@{text "!"} - & @{text "\"} & & @{text "\"} & & @{text "\"} \\ - @{attribute elim} @{attribute intro} - & @{text "\"} & & @{text "\"} & & @{text "\"} \\ - @{attribute iff} - & @{text "\"} & & @{text "\"} & @{text "\"} & @{text "\"} \\ - @{attribute iff}@{text "?"} - & @{text "\"} \\ - @{attribute elim}@{text "?"} @{attribute intro}@{text "?"} - & @{text "\"} \\ - @{attribute simp} - & & & & @{text "\"} & @{text "\"} \\ - @{attribute cong} - & & & & @{text "\"} & @{text "\"} \\ - @{attribute split} - & & & & @{text "\"} & @{text "\"} \\ - \end{tabular} -*} - - -section {* Emulating tactic scripts *} - -subsection {* Commands *} - -text {* - \begin{tabular}{ll} - @{command "apply"}~@{text m} & apply proof method at initial position \\ - @{command "apply_end"}~@{text m} & apply proof method near terminal position \\ - @{command "done"} & complete proof \\ - @{command "defer"}~@{text n} & move subgoal to end \\ - @{command "prefer"}~@{text n} & move subgoal to beginning \\ - @{command "back"} & backtrack last command \\ - \end{tabular} -*} - - -subsection {* Methods *} - -text {* - \begin{tabular}{ll} - @{method rule_tac}~@{text insts} & resolution (with instantiation) \\ - @{method erule_tac}~@{text insts} & elim-resolution (with instantiation) \\ - @{method drule_tac}~@{text insts} & destruct-resolution (with instantiation) \\ - @{method frule_tac}~@{text insts} & forward-resolution (with instantiation) \\ - @{method cut_tac}~@{text insts} & insert facts (with instantiation) \\ - @{method thin_tac}~@{text \} & delete assumptions \\ - @{method subgoal_tac}~@{text \} & new claims \\ - @{method rename_tac}~@{text x} & rename innermost goal parameters \\ - @{method rotate_tac}~@{text n} & rotate assumptions of goal \\ - @{method tactic}~@{text "text"} & arbitrary ML tactic \\ - @{method case_tac}~@{text t} & exhaustion (datatypes) \\ - @{method induct_tac}~@{text x} & induction (datatypes) \\ - @{method ind_cases}~@{text t} & exhaustion + simplification (inductive predicates) \\ - \end{tabular} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Spec.thy --- a/doc-src/IsarRef/Thy/Spec.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1337 +0,0 @@ -theory Spec -imports Base Main -begin - -chapter {* Specifications *} - -text {* - The Isabelle/Isar theory format integrates specifications and - proofs, supporting interactive development with unlimited undo - operation. There is an integrated document preparation system (see - \chref{ch:document-prep}), for typesetting formal developments - together with informal text. The resulting hyper-linked PDF - documents can be used both for WWW presentation and printed copies. - - The Isar proof language (see \chref{ch:proofs}) is embedded into the - theory language as a proper sub-language. Proof mode is entered by - stating some @{command theorem} or @{command lemma} at the theory - level, and left again with the final conclusion (e.g.\ via @{command - qed}). Some theory specification mechanisms also require a proof, - such as @{command typedef} in HOL, which demands non-emptiness of - the representing sets. -*} - - -section {* Defining theories \label{sec:begin-thy} *} - -text {* - \begin{matharray}{rcl} - @{command_def "theory"} & : & @{text "toplevel \ theory"} \\ - @{command_def (global) "end"} & : & @{text "theory \ toplevel"} \\ - \end{matharray} - - Isabelle/Isar theories are defined via theory files, which may - contain both specifications and proofs; occasionally definitional - mechanisms also require some explicit proof. The theory body may be - sub-structured by means of \emph{local theory targets}, such as - @{command "locale"} and @{command "class"}. - - The first proper command of a theory is @{command "theory"}, which - indicates imports of previous theories and optional dependencies on - other source files (usually in ML). Just preceding the initial - @{command "theory"} command there may be an optional @{command - "header"} declaration, which is only relevant to document - preparation: see also the other section markup commands in - \secref{sec:markup}. - - A theory is concluded by a final @{command (global) "end"} command, - one that does not belong to a local theory target. No further - commands may follow such a global @{command (global) "end"}, - although some user-interfaces might pretend that trailing input is - admissible. - - @{rail " - @@{command theory} @{syntax name} imports \\ keywords? uses? @'begin' - ; - imports: @'imports' (@{syntax name} +) - ; - keywords: @'keywords' ((@{syntax string} +) ('::' @{syntax name} @{syntax tags})? + @'and') - ; - uses: @'uses' ((@{syntax name} | @{syntax parname}) +) - "} - - \begin{description} - - \item @{command "theory"}~@{text "A \ B\<^sub>1 \ B\<^sub>n \"} - starts a new theory @{text A} based on the merge of existing - theories @{text "B\<^sub>1 \ B\<^sub>n"}. Due to the possibility to import more - than one ancestor, the resulting theory structure of an Isabelle - session forms a directed acyclic graph (DAG). Isabelle takes care - that sources contributing to the development graph are always - up-to-date: changed files are automatically rechecked whenever a - theory header specification is processed. - - The optional @{keyword_def "keywords"} specification declares outer - syntax (\chref{ch:outer-syntax}) that is introduced in this theory - later on (rare in end-user applications). Both minor keywords and - major keywords of the Isar command language need to be specified, in - order to make parsing of proof documents work properly. Command - keywords need to be classified according to their structural role in - the formal text. Examples may be seen in Isabelle/HOL sources - itself, such as @{keyword "keywords"}~@{verbatim "\"typedef\""} - @{text ":: thy_goal"} or @{keyword "keywords"}~@{verbatim - "\"datatype\""} @{text ":: thy_decl"} for theory-level declarations - with and without proof, respectively. Additional @{syntax tags} - provide defaults for document preparation (\secref{sec:tags}). - - The optional @{keyword_def "uses"} specification declares additional - dependencies on external files (notably ML sources). Files will be - loaded immediately (as ML), unless the name is parenthesized. The - latter case records a dependency that needs to be resolved later in - the text, usually via explicit @{command_ref "use"} for ML files; - other file formats require specific load commands defined by the - corresponding tools or packages. - - \item @{command (global) "end"} concludes the current theory - definition. Note that some other commands, e.g.\ local theory - targets @{command locale} or @{command class} may involve a - @{keyword "begin"} that needs to be matched by @{command (local) - "end"}, according to the usual rules for nested blocks. - - \end{description} -*} - - -section {* Local theory targets \label{sec:target} *} - -text {* - \begin{matharray}{rcll} - @{command_def "context"} & : & @{text "theory \ local_theory"} \\ - @{command_def (local) "end"} & : & @{text "local_theory \ theory"} \\ - \end{matharray} - - A local theory target is a context managed separately within the - enclosing theory. Contexts may introduce parameters (fixed - variables) and assumptions (hypotheses). Definitions and theorems - depending on the context may be added incrementally later on. - - \emph{Named contexts} refer to locales (cf.\ \secref{sec:locale}) or - type classes (cf.\ \secref{sec:class}); the name ``@{text "-"}'' - signifies the global theory context. - - \emph{Unnamed contexts} may introduce additional parameters and - assumptions, and results produced in the context are generalized - accordingly. Such auxiliary contexts may be nested within other - targets, like @{command "locale"}, @{command "class"}, @{command - "instantiation"}, @{command "overloading"}. - - @{rail " - @@{command context} @{syntax nameref} @'begin' - ; - @@{command context} @{syntax_ref \"includes\"}? (@{syntax context_elem} * ) @'begin' - ; - @{syntax_def target}: '(' @'in' @{syntax nameref} ')' - "} - - \begin{description} - - \item @{command "context"}~@{text "c \"} opens a named - context, by recommencing an existing locale or class @{text c}. - Note that locale and class definitions allow to include the - @{keyword "begin"} keyword as well, in order to continue the local - theory immediately after the initial specification. - - \item @{command "context"}~@{text "bundles elements \"} opens - an unnamed context, by extending the enclosing global or local - theory target by the given declaration bundles (\secref{sec:bundle}) - and context elements (@{text "\"}, @{text "\"} - etc.). This means any results stemming from definitions and proofs - in the extended context will be exported into the enclosing target - by lifting over extra parameters and premises. - - \item @{command (local) "end"} concludes the current local theory, - according to the nesting of contexts. Note that a global @{command - (global) "end"} has a different meaning: it concludes the theory - itself (\secref{sec:begin-thy}). - - \item @{text "("}@{keyword_def "in"}~@{text "c)"} given after any - local theory command specifies an immediate target, e.g.\ - ``@{command "definition"}~@{text "(\ c) \"}'' or ``@{command - "theorem"}~@{text "(\ c) \"}''. This works both in a local or - global theory context; the current target context will be suspended - for this command only. Note that ``@{text "(\ -)"}'' will - always produce a global result independently of the current target - context. - - \end{description} - - The exact meaning of results produced within a local theory context - depends on the underlying target infrastructure (locale, type class - etc.). The general idea is as follows, considering a context named - @{text c} with parameter @{text x} and assumption @{text "A[x]"}. - - Definitions are exported by introducing a global version with - additional arguments; a syntactic abbreviation links the long form - with the abstract version of the target context. For example, - @{text "a \ t[x]"} becomes @{text "c.a ?x \ t[?x]"} at the theory - level (for arbitrary @{text "?x"}), together with a local - abbreviation @{text "c \ c.a x"} in the target context (for the - fixed parameter @{text x}). - - Theorems are exported by discharging the assumptions and - generalizing the parameters of the context. For example, @{text "a: - B[x]"} becomes @{text "c.a: A[?x] \ B[?x]"}, again for arbitrary - @{text "?x"}. - - \medskip The Isabelle/HOL library contains numerous applications of - locales and classes, e.g.\ see @{file "~~/src/HOL/Algebra"}. An - example for an unnamed auxiliary contexts is given in @{file - "~~/src/HOL/Isar_Examples/Group_Context.thy"}. *} - - -section {* Bundled declarations \label{sec:bundle} *} - -text {* - \begin{matharray}{rcl} - @{command_def "bundle"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "print_bundles"}@{text "\<^sup>*"} & : & @{text "context \ "} \\ - @{command_def "include"} & : & @{text "proof(state) \ proof(state)"} \\ - @{command_def "including"} & : & @{text "proof(prove) \ proof(prove)"} \\ - @{keyword_def "includes"} & : & syntax \\ - \end{matharray} - - The outer syntax of fact expressions (\secref{sec:syn-att}) involves - theorems and attributes, which are evaluated in the context and - applied to it. Attributes may declare theorems to the context, as - in @{text "this_rule [intro] that_rule [elim]"} for example. - Configuration options (\secref{sec:config}) are special declaration - attributes that operate on the context without a theorem, as in - @{text "[[show_types = false]]"} for example. - - Expressions of this form may be defined as \emph{bundled - declarations} in the context, and included in other situations later - on. Including declaration bundles augments a local context casually - without logical dependencies, which is in contrast to locales and - locale interpretation (\secref{sec:locale}). - - @{rail " - @@{command bundle} @{syntax target}? \\ - @{syntax name} '=' @{syntax thmrefs} (@'for' (@{syntax vars} + @'and'))? - ; - (@@{command include} | @@{command including}) (@{syntax nameref}+) - ; - @{syntax_def \"includes\"}: @'includes' (@{syntax nameref}+) - "} - - \begin{description} - - \item @{command bundle}~@{text "b = decls"} defines a bundle of - declarations in the current context. The RHS is similar to the one - of the @{command declare} command. Bundles defined in local theory - targets are subject to transformations via morphisms, when moved - into different application contexts; this works analogously to any - other local theory specification. - - \item @{command print_bundles} prints the named bundles that are - available in the current context. - - \item @{command include}~@{text "b\<^sub>1 \ b\<^sub>n"} includes the declarations - from the given bundles into the current proof body context. This is - analogous to @{command "note"} (\secref{sec:proof-facts}) with the - expanded bundles. - - \item @{command including} is similar to @{command include}, but - works in proof refinement (backward mode). This is analogous to - @{command "using"} (\secref{sec:proof-facts}) with the expanded - bundles. - - \item @{keyword "includes"}~@{text "b\<^sub>1 \ b\<^sub>n"} is similar to - @{command include}, but works in situations where a specification - context is constructed, notably for @{command context} and long - statements of @{command theorem} etc. - - \end{description} - - Here is an artificial example of bundling various configuration - options: *} - -bundle trace = [[simp_trace, blast_trace, linarith_trace, metis_trace, smt_trace]] - -lemma "x = x" - including trace by metis - - -section {* Basic specification elements *} - -text {* - \begin{matharray}{rcll} - @{command_def "axiomatization"} & : & @{text "theory \ theory"} & (axiomatic!) \\ - @{command_def "definition"} & : & @{text "local_theory \ local_theory"} \\ - @{attribute_def "defn"} & : & @{text attribute} \\ - @{command_def "abbreviation"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "print_abbrevs"}@{text "\<^sup>*"} & : & @{text "context \ "} \\ - \end{matharray} - - These specification mechanisms provide a slightly more abstract view - than the underlying primitives of @{command "consts"}, @{command - "defs"} (see \secref{sec:consts}), and @{command "axioms"} (see - \secref{sec:axms-thms}). In particular, type-inference is commonly - available, and result names need not be given. - - @{rail " - @@{command axiomatization} @{syntax \"fixes\"}? (@'where' specs)? - ; - @@{command definition} @{syntax target}? \\ - (decl @'where')? @{syntax thmdecl}? @{syntax prop} - ; - @@{command abbreviation} @{syntax target}? @{syntax mode}? \\ - (decl @'where')? @{syntax prop} - ; - - @{syntax_def \"fixes\"}: ((@{syntax name} ('::' @{syntax type})? - @{syntax mixfix}? | @{syntax vars}) + @'and') - ; - specs: (@{syntax thmdecl}? @{syntax props} + @'and') - ; - decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? - "} - - \begin{description} - - \item @{command "axiomatization"}~@{text "c\<^sub>1 \ c\<^sub>m \ \\<^sub>1 \ \\<^sub>n"} - introduces several constants simultaneously and states axiomatic - properties for these. The constants are marked as being specified - once and for all, which prevents additional specifications being - issued later on. - - Note that axiomatic specifications are only appropriate when - declaring a new logical system; axiomatic specifications are - restricted to global theory contexts. Normal applications should - only use definitional mechanisms! - - \item @{command "definition"}~@{text "c \ eq"} produces an - internal definition @{text "c \ t"} according to the specification - given as @{text eq}, which is then turned into a proven fact. The - given proposition may deviate from internal meta-level equality - according to the rewrite rules declared as @{attribute defn} by the - object-logic. This usually covers object-level equality @{text "x = - y"} and equivalence @{text "A \ B"}. End-users normally need not - change the @{attribute defn} setup. - - Definitions may be presented with explicit arguments on the LHS, as - well as additional conditions, e.g.\ @{text "f x y = t"} instead of - @{text "f \ \x y. t"} and @{text "y \ 0 \ g x y = u"} instead of an - unrestricted @{text "g \ \x y. u"}. - - \item @{command "abbreviation"}~@{text "c \ eq"} introduces a - syntactic constant which is associated with a certain term according - to the meta-level equality @{text eq}. - - Abbreviations participate in the usual type-inference process, but - are expanded before the logic ever sees them. Pretty printing of - terms involves higher-order rewriting with rules stemming from - reverted abbreviations. This needs some care to avoid overlapping - or looping syntactic replacements! - - The optional @{text mode} specification restricts output to a - particular print mode; using ``@{text input}'' here achieves the - effect of one-way abbreviations. The mode may also include an - ``@{keyword "output"}'' qualifier that affects the concrete syntax - declared for abbreviations, cf.\ @{command "syntax"} in - \secref{sec:syn-trans}. - - \item @{command "print_abbrevs"} prints all constant abbreviations - of the current context. - - \end{description} -*} - - -section {* Generic declarations *} - -text {* - \begin{matharray}{rcl} - @{command_def "declaration"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "syntax_declaration"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "declare"} & : & @{text "local_theory \ local_theory"} \\ - \end{matharray} - - Arbitrary operations on the background context may be wrapped-up as - generic declaration elements. Since the underlying concept of local - theories may be subject to later re-interpretation, there is an - additional dependency on a morphism that tells the difference of the - original declaration context wrt.\ the application context - encountered later on. A fact declaration is an important special - case: it consists of a theorem which is applied to the context by - means of an attribute. - - @{rail " - (@@{command declaration} | @@{command syntax_declaration}) - ('(' @'pervasive' ')')? \\ @{syntax target}? @{syntax text} - ; - @@{command declare} @{syntax target}? (@{syntax thmrefs} + @'and') - "} - - \begin{description} - - \item @{command "declaration"}~@{text d} adds the declaration - function @{text d} of ML type @{ML_type declaration}, to the current - local theory under construction. In later application contexts, the - function is transformed according to the morphisms being involved in - the interpretation hierarchy. - - If the @{text "(pervasive)"} option is given, the corresponding - declaration is applied to all possible contexts involved, including - the global background theory. - - \item @{command "syntax_declaration"} is similar to @{command - "declaration"}, but is meant to affect only ``syntactic'' tools by - convention (such as notation and type-checking information). - - \item @{command "declare"}~@{text thms} declares theorems to the - current local theory context. No theorem binding is involved here, - unlike @{command "theorems"} or @{command "lemmas"} (cf.\ - \secref{sec:axms-thms}), so @{command "declare"} only has the effect - of applying attributes as included in the theorem specification. - - \end{description} -*} - - -section {* Locales \label{sec:locale} *} - -text {* - Locales are parametric named local contexts, consisting of a list of - declaration elements that are modeled after the Isar proof context - commands (cf.\ \secref{sec:proof-context}). -*} - - -subsection {* Locale expressions \label{sec:locale-expr} *} - -text {* - A \emph{locale expression} denotes a structured context composed of - instances of existing locales. The context consists of a list of - instances of declaration elements from the locales. Two locale - instances are equal if they are of the same locale and the - parameters are instantiated with equivalent terms. Declaration - elements from equal instances are never repeated, thus avoiding - duplicate declarations. More precisely, declarations from instances - that are subsumed by earlier instances are omitted. - - @{rail " - @{syntax_def locale_expr}: (instance + '+') (@'for' (@{syntax \"fixes\"} + @'and'))? - ; - instance: (qualifier ':')? @{syntax nameref} (pos_insts | named_insts) - ; - qualifier: @{syntax name} ('?' | '!')? - ; - pos_insts: ('_' | @{syntax term})* - ; - named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and') - "} - - A locale instance consists of a reference to a locale and either - positional or named parameter instantiations. Identical - instantiations (that is, those that instante a parameter by itself) - may be omitted. The notation `@{text "_"}' enables to omit the - instantiation for a parameter inside a positional instantiation. - - Terms in instantiations are from the context the locale expressions - is declared in. Local names may be added to this context with the - optional @{keyword "for"} clause. This is useful for shadowing names - bound in outer contexts, and for declaring syntax. In addition, - syntax declarations from one instance are effective when parsing - subsequent instances of the same expression. - - Instances have an optional qualifier which applies to names in - declarations. Names include local definitions and theorem names. - If present, the qualifier itself is either optional - (``\texttt{?}''), which means that it may be omitted on input of the - qualified name, or mandatory (``\texttt{!}''). If neither - ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default - is used. For @{command "interpretation"} and @{command "interpret"} - the default is ``mandatory'', for @{command "locale"} and @{command - "sublocale"} the default is ``optional''. -*} - - -subsection {* Locale declarations *} - -text {* - \begin{matharray}{rcl} - @{command_def "locale"} & : & @{text "theory \ local_theory"} \\ - @{command_def "print_locale"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_locales"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{method_def intro_locales} & : & @{text method} \\ - @{method_def unfold_locales} & : & @{text method} \\ - \end{matharray} - - \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes} - \indexisarelem{defines}\indexisarelem{notes} - @{rail " - @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'? - ; - @@{command print_locale} '!'? @{syntax nameref} - ; - @{syntax_def locale}: @{syntax context_elem}+ | - @{syntax locale_expr} ('+' (@{syntax context_elem}+))? - ; - @{syntax_def context_elem}: - @'fixes' (@{syntax \"fixes\"} + @'and') | - @'constrains' (@{syntax name} '::' @{syntax type} + @'and') | - @'assumes' (@{syntax props} + @'and') | - @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') | - @'notes' (@{syntax thmdef}? @{syntax thmrefs} + @'and') - "} - - \begin{description} - - \item @{command "locale"}~@{text "loc = import + body"} defines a - new locale @{text loc} as a context consisting of a certain view of - existing locales (@{text import}) plus some additional elements - (@{text body}). Both @{text import} and @{text body} are optional; - the degenerate form @{command "locale"}~@{text loc} defines an empty - locale, which may still be useful to collect declarations of facts - later on. Type-inference on locale expressions automatically takes - care of the most general typing that the combined context elements - may acquire. - - The @{text import} consists of a structured locale expression; see - \secref{sec:proof-context} above. Its for clause defines the local - parameters of the @{text import}. In addition, locale parameters - whose instantance is omitted automatically extend the (possibly - empty) for clause: they are inserted at its beginning. This means - that these parameters may be referred to from within the expression - and also in the subsequent context elements and provides a - notational convenience for the inheritance of parameters in locale - declarations. - - The @{text body} consists of context elements. - - \begin{description} - - \item @{element "fixes"}~@{text "x :: \ (mx)"} declares a local - parameter of type @{text \} and mixfix annotation @{text mx} (both - are optional). The special syntax declaration ``@{text - "(\)"}'' means that @{text x} may be referenced - implicitly in this context. - - \item @{element "constrains"}~@{text "x :: \"} introduces a type - constraint @{text \} on the local parameter @{text x}. This - element is deprecated. The type constraint should be introduced in - the for clause or the relevant @{element "fixes"} element. - - \item @{element "assumes"}~@{text "a: \\<^sub>1 \ \\<^sub>n"} - introduces local premises, similar to @{command "assume"} within a - proof (cf.\ \secref{sec:proof-context}). - - \item @{element "defines"}~@{text "a: x \ t"} defines a previously - declared parameter. This is similar to @{command "def"} within a - proof (cf.\ \secref{sec:proof-context}), but @{element "defines"} - takes an equational proposition instead of variable-term pair. The - left-hand side of the equation may have additional arguments, e.g.\ - ``@{element "defines"}~@{text "f x\<^sub>1 \ x\<^sub>n \ t"}''. - - \item @{element "notes"}~@{text "a = b\<^sub>1 \ b\<^sub>n"} - reconsiders facts within a local context. Most notably, this may - include arbitrary declarations in any attribute specifications - included here, e.g.\ a local @{attribute simp} rule. - - The initial @{text import} specification of a locale expression - maintains a dynamic relation to the locales being referenced - (benefiting from any later fact declarations in the obvious manner). - - \end{description} - - Note that ``@{text "(\ p\<^sub>1 \ p\<^sub>n)"}'' patterns given - in the syntax of @{element "assumes"} and @{element "defines"} above - are illegal in locale definitions. In the long goal format of - \secref{sec:goals}, term bindings may be included as expected, - though. - - \medskip Locale specifications are ``closed up'' by - turning the given text into a predicate definition @{text - loc_axioms} and deriving the original assumptions as local lemmas - (modulo local definitions). The predicate statement covers only the - newly specified assumptions, omitting the content of included locale - expressions. The full cumulative view is only provided on export, - involving another predicate @{text loc} that refers to the complete - specification text. - - In any case, the predicate arguments are those locale parameters - that actually occur in the respective piece of text. Also note that - these predicates operate at the meta-level in theory, but the locale - packages attempts to internalize statements according to the - object-logic setup (e.g.\ replacing @{text \} by @{text \}, and - @{text "\"} by @{text "\"} in HOL; see also - \secref{sec:object-logic}). Separate introduction rules @{text - loc_axioms.intro} and @{text loc.intro} are provided as well. - - \item @{command "print_locale"}~@{text "locale"} prints the - contents of the named locale. The command omits @{element "notes"} - elements by default. Use @{command "print_locale"}@{text "!"} to - have them included. - - \item @{command "print_locales"} prints the names of all locales - of the current theory. - - \item @{method intro_locales} and @{method unfold_locales} - repeatedly expand all introduction rules of locale predicates of the - theory. While @{method intro_locales} only applies the @{text - loc.intro} introduction rules and therefore does not decend to - assumptions, @{method unfold_locales} is more aggressive and applies - @{text loc_axioms.intro} as well. Both methods are aware of locale - specifications entailed by the context, both from target statements, - and from interpretations (see below). New goals that are entailed - by the current context are discharged automatically. - - \end{description} -*} - - -subsection {* Locale interpretation *} - -text {* - \begin{matharray}{rcl} - @{command_def "interpretation"} & : & @{text "theory \ proof(prove)"} \\ - @{command_def "interpret"} & : & @{text "proof(state) | proof(chain) \ proof(prove)"} \\ - @{command_def "sublocale"} & : & @{text "theory \ proof(prove)"} \\ - @{command_def "print_dependencies"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "print_interps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - \end{matharray} - - Locale expressions may be instantiated, and the instantiated facts - added to the current context. This requires a proof of the - instantiated specification and is called \emph{locale - interpretation}. Interpretation is possible in locales (command - @{command "sublocale"}), theories (command @{command - "interpretation"}) and also within proof bodies (command @{command - "interpret"}). - - @{rail " - @@{command interpretation} @{syntax locale_expr} equations? - ; - @@{command interpret} @{syntax locale_expr} equations? - ; - @@{command sublocale} @{syntax nameref} ('<' | '\') @{syntax locale_expr} \\ - equations? - ; - @@{command print_dependencies} '!'? @{syntax locale_expr} - ; - @@{command print_interps} @{syntax nameref} - ; - - equations: @'where' (@{syntax thmdecl}? @{syntax prop} + @'and') - "} - - \begin{description} - - \item @{command "interpretation"}~@{text "expr \ eqns"} - interprets @{text expr} in the theory. The command generates proof - obligations for the instantiated specifications (assumes and defines - elements). Once these are discharged by the user, instantiated - facts are added to the theory in a post-processing phase. - - Free variables in the interpreted expression are allowed. They are - turned into schematic variables in the generated declarations. In - order to use a free variable whose name is already bound in the - context --- for example, because a constant of that name exists --- - it may be added to the @{keyword "for"} clause. - - Additional equations, which are unfolded during - post-processing, may be given after the keyword @{keyword "where"}. - This is useful for interpreting concepts introduced through - definitions. The equations must be proved. - - The command is aware of interpretations already active in the - theory, but does not simplify the goal automatically. In order to - simplify the proof obligations use methods @{method intro_locales} - or @{method unfold_locales}. Post-processing is not applied to - facts of interpretations that are already active. This avoids - duplication of interpreted facts, in particular. Note that, in the - case of a locale with import, parts of the interpretation may - already be active. The command will only process facts for new - parts. - - Adding facts to locales has the effect of adding interpreted facts - to the theory for all interpretations as well. That is, - interpretations dynamically participate in any facts added to - locales. Note that if a theory inherits additional facts for a - locale through one parent and an interpretation of that locale - through another parent, the additional facts will not be - interpreted. - - \item @{command "interpret"}~@{text "expr \ eqns"} interprets - @{text expr} in the proof context and is otherwise similar to - interpretation in theories. Note that rewrite rules given to - @{command "interpret"} after the @{keyword "where"} keyword should be - explicitly universally quantified. - - \item @{command "sublocale"}~@{text "name \ expr \ - eqns"} - interprets @{text expr} in the locale @{text name}. A proof that - the specification of @{text name} implies the specification of - @{text expr} is required. As in the localized version of the - theorem command, the proof is in the context of @{text name}. After - the proof obligation has been discharged, the facts of @{text expr} - become part of locale @{text name} as \emph{derived} context - elements and are available when the context @{text name} is - subsequently entered. Note that, like import, this is dynamic: - facts added to a locale part of @{text expr} after interpretation - become also available in @{text name}. - - Only specification fragments of @{text expr} that are not already - part of @{text name} (be it imported, derived or a derived fragment - of the import) are considered in this process. This enables - circular interpretations provided that no infinite chains are - generated in the locale hierarchy. - - If interpretations of @{text name} exist in the current theory, the - command adds interpretations for @{text expr} as well, with the same - qualifier, although only for fragments of @{text expr} that are not - interpreted in the theory already. - - Equations given after @{keyword "where"} amend the morphism through - which @{text expr} is interpreted. This enables to map definitions - from the interpreted locales to entities of @{text name}. This - feature is experimental. - - \item @{command "print_dependencies"}~@{text "expr"} is useful for - understanding the effect of an interpretation of @{text "expr"}. It - lists all locale instances for which interpretations would be added - to the current context. Variant @{command - "print_dependencies"}@{text "!"} prints all locale instances that - would be considered for interpretation, and would be interpreted in - an empty context (that is, without interpretations). - - \item @{command "print_interps"}~@{text "locale"} lists all - interpretations of @{text "locale"} in the current theory or proof - context, including those due to a combination of a @{command - "interpretation"} or @{command "interpret"} and one or several - @{command "sublocale"} declarations. - - \end{description} - - \begin{warn} - Since attributes are applied to interpreted theorems, - interpretation may modify the context of common proof tools, e.g.\ - the Simplifier or Classical Reasoner. As the behavior of such - tools is \emph{not} stable under interpretation morphisms, manual - declarations might have to be added to the target context of the - interpretation to revert such declarations. - \end{warn} - - \begin{warn} - An interpretation in a theory or proof context may subsume previous - interpretations. This happens if the same specification fragment - is interpreted twice and the instantiation of the second - interpretation is more general than the interpretation of the - first. The locale package does not attempt to remove subsumed - interpretations. - \end{warn} -*} - - -section {* Classes \label{sec:class} *} - -text {* - \begin{matharray}{rcl} - @{command_def "class"} & : & @{text "theory \ local_theory"} \\ - @{command_def "instantiation"} & : & @{text "theory \ local_theory"} \\ - @{command_def "instance"} & : & @{text "local_theory \ local_theory"} \\ - @{command "instance"} & : & @{text "theory \ proof(prove)"} \\ - @{command_def "subclass"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "print_classes"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{command_def "class_deps"}@{text "\<^sup>*"} & : & @{text "context \"} \\ - @{method_def intro_classes} & : & @{text method} \\ - \end{matharray} - - A class is a particular locale with \emph{exactly one} type variable - @{text \}. Beyond the underlying locale, a corresponding type class - is established which is interpreted logically as axiomatic type - class \cite{Wenzel:1997:TPHOL} whose logical content are the - assumptions of the locale. Thus, classes provide the full - generality of locales combined with the commodity of type classes - (notably type-inference). See \cite{isabelle-classes} for a short - tutorial. - - @{rail " - @@{command class} class_spec @'begin'? - ; - class_spec: @{syntax name} '=' - ((@{syntax nameref} '+' (@{syntax context_elem}+)) | - @{syntax nameref} | (@{syntax context_elem}+)) - ; - @@{command instantiation} (@{syntax nameref} + @'and') '::' @{syntax arity} @'begin' - ; - @@{command instance} (() | (@{syntax nameref} + @'and') '::' @{syntax arity} | - @{syntax nameref} ('<' | '\') @{syntax nameref} ) - ; - @@{command subclass} @{syntax target}? @{syntax nameref} - "} - - \begin{description} - - \item @{command "class"}~@{text "c = superclasses + body"} defines - a new class @{text c}, inheriting from @{text superclasses}. This - introduces a locale @{text c} with import of all locales @{text - superclasses}. - - Any @{element "fixes"} in @{text body} are lifted to the global - theory level (\emph{class operations} @{text "f\<^sub>1, \, - f\<^sub>n"} of class @{text c}), mapping the local type parameter - @{text \} to a schematic type variable @{text "?\ :: c"}. - - Likewise, @{element "assumes"} in @{text body} are also lifted, - mapping each local parameter @{text "f :: \[\]"} to its - corresponding global constant @{text "f :: \[?\ :: c]"}. The - corresponding introduction rule is provided as @{text - c_class_axioms.intro}. This rule should be rarely needed directly - --- the @{method intro_classes} method takes care of the details of - class membership proofs. - - \item @{command "instantiation"}~@{text "t :: (s\<^sub>1, \, s\<^sub>n)s - \"} opens a theory target (cf.\ \secref{sec:target}) which - allows to specify class operations @{text "f\<^sub>1, \, f\<^sub>n"} corresponding - to sort @{text s} at the particular type instance @{text "(\\<^sub>1 :: s\<^sub>1, - \, \\<^sub>n :: s\<^sub>n) t"}. A plain @{command "instance"} command in the - target body poses a goal stating these type arities. The target is - concluded by an @{command_ref (local) "end"} command. - - Note that a list of simultaneous type constructors may be given; - this corresponds nicely to mutually recursive type definitions, e.g.\ - in Isabelle/HOL. - - \item @{command "instance"} in an instantiation target body sets - up a goal stating the type arities claimed at the opening @{command - "instantiation"}. The proof would usually proceed by @{method - intro_classes}, and then establish the characteristic theorems of - the type classes involved. After finishing the proof, the - background theory will be augmented by the proven type arities. - - On the theory level, @{command "instance"}~@{text "t :: (s\<^sub>1, \, - s\<^sub>n)s"} provides a convenient way to instantiate a type class with no - need to specify operations: one can continue with the - instantiation proof immediately. - - \item @{command "subclass"}~@{text c} in a class context for class - @{text d} sets up a goal stating that class @{text c} is logically - contained in class @{text d}. After finishing the proof, class - @{text d} is proven to be subclass @{text c} and the locale @{text - c} is interpreted into @{text d} simultaneously. - - A weakend form of this is available through a further variant of - @{command instance}: @{command instance}~@{text "c\<^sub>1 \ c\<^sub>2"} opens - a proof that class @{text "c\<^isub>2"} implies @{text "c\<^isub>1"} without reference - to the underlying locales; this is useful if the properties to prove - the logical connection are not sufficent on the locale level but on - the theory level. - - \item @{command "print_classes"} prints all classes in the current - theory. - - \item @{command "class_deps"} visualizes all classes and their - subclass relations as a Hasse diagram. - - \item @{method intro_classes} repeatedly expands all class - introduction rules of this theory. Note that this method usually - needs not be named explicitly, as it is already included in the - default proof step (e.g.\ of @{command "proof"}). In particular, - instantiation of trivial (syntactic) classes may be performed by a - single ``@{command ".."}'' proof step. - - \end{description} -*} - - -subsection {* The class target *} - -text {* - %FIXME check - - A named context may refer to a locale (cf.\ \secref{sec:target}). - If this locale is also a class @{text c}, apart from the common - locale target behaviour the following happens. - - \begin{itemize} - - \item Local constant declarations @{text "g[\]"} referring to the - local type parameter @{text \} and local parameters @{text "f[\]"} - are accompanied by theory-level constants @{text "g[?\ :: c]"} - referring to theory-level class operations @{text "f[?\ :: c]"}. - - \item Local theorem bindings are lifted as are assumptions. - - \item Local syntax refers to local operations @{text "g[\]"} and - global operations @{text "g[?\ :: c]"} uniformly. Type inference - resolves ambiguities. In rare cases, manual type annotations are - needed. - - \end{itemize} -*} - - -subsection {* Co-regularity of type classes and arities *} - -text {* The class relation together with the collection of - type-constructor arities must obey the principle of - \emph{co-regularity} as defined below. - - \medskip For the subsequent formulation of co-regularity we assume - that the class relation is closed by transitivity and reflexivity. - Moreover the collection of arities @{text "t :: (\<^vec>s)c"} is - completed such that @{text "t :: (\<^vec>s)c"} and @{text "c \ c'"} - implies @{text "t :: (\<^vec>s)c'"} for all such declarations. - - Treating sorts as finite sets of classes (meaning the intersection), - the class relation @{text "c\<^sub>1 \ c\<^sub>2"} is extended to sorts as - follows: - \[ - @{text "s\<^sub>1 \ s\<^sub>2 \ \c\<^sub>2 \ s\<^sub>2. \c\<^sub>1 \ s\<^sub>1. c\<^sub>1 \ c\<^sub>2"} - \] - - This relation on sorts is further extended to tuples of sorts (of - the same length) in the component-wise way. - - \smallskip Co-regularity of the class relation together with the - arities relation means: - \[ - @{text "t :: (\<^vec>s\<^sub>1)c\<^sub>1 \ t :: (\<^vec>s\<^sub>2)c\<^sub>2 \ c\<^sub>1 \ c\<^sub>2 \ \<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2"} - \] - \noindent for all such arities. In other words, whenever the result - classes of some type-constructor arities are related, then the - argument sorts need to be related in the same way. - - \medskip Co-regularity is a very fundamental property of the - order-sorted algebra of types. For example, it entails principle - types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}. -*} - - -section {* Unrestricted overloading *} - -text {* - \begin{matharray}{rcl} - @{command_def "overloading"} & : & @{text "theory \ local_theory"} \\ - \end{matharray} - - Isabelle/Pure's definitional schemes support certain forms of - overloading (see \secref{sec:consts}). Overloading means that a - constant being declared as @{text "c :: \ decl"} may be - defined separately on type instances - @{text "c :: (\\<^sub>1, \, \\<^sub>n) t decl"} - for each type constructor @{text t}. At most occassions - overloading will be used in a Haskell-like fashion together with - type classes by means of @{command "instantiation"} (see - \secref{sec:class}). Sometimes low-level overloading is desirable. - The @{command "overloading"} target provides a convenient view for - end-users. - - @{rail " - @@{command overloading} ( spec + ) @'begin' - ; - spec: @{syntax name} ( '==' | '\' ) @{syntax term} ( '(' @'unchecked' ')' )? - "} - - \begin{description} - - \item @{command "overloading"}~@{text "x\<^sub>1 \ c\<^sub>1 :: \\<^sub>1 \ \ x\<^sub>n \ c\<^sub>n :: \\<^sub>n \"} - opens a theory target (cf.\ \secref{sec:target}) which allows to - specify constants with overloaded definitions. These are identified - by an explicitly given mapping from variable names @{text "x\<^sub>i"} to - constants @{text "c\<^sub>i"} at particular type instances. The - definitions themselves are established using common specification - tools, using the names @{text "x\<^sub>i"} as reference to the - corresponding constants. The target is concluded by @{command - (local) "end"}. - - A @{text "(unchecked)"} option disables global dependency checks for - the corresponding definition, which is occasionally useful for - exotic overloading (see \secref{sec:consts} for a precise description). - It is at the discretion of the user to avoid - malformed theory specifications! - - \end{description} -*} - - -section {* Incorporating ML code \label{sec:ML} *} - -text {* - \begin{matharray}{rcl} - @{command_def "ML_file"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "ML"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "ML_prf"} & : & @{text "proof \ proof"} \\ - @{command_def "ML_val"} & : & @{text "any \"} \\ - @{command_def "ML_command"} & : & @{text "any \"} \\ - @{command_def "setup"} & : & @{text "theory \ theory"} \\ - @{command_def "local_setup"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "attribute_setup"} & : & @{text "theory \ theory"} \\ - \end{matharray} - - @{rail " - @@{command ML_file} @{syntax name} - ; - (@@{command ML} | @@{command ML_prf} | @@{command ML_val} | - @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text} - ; - @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}? - "} - - \begin{description} - - \item @{command "ML_file"}~@{text "name"} reads and evaluates the - given ML file. The current theory context is passed down to the ML - toplevel and may be modified, using @{ML "Context.>>"} or derived ML - commands. Top-level ML bindings are stored within the (global or - local) theory context. - - \item @{command "ML"}~@{text "text"} is similar to @{command - "ML_file"}, but evaluates directly the given @{text "text"}. - Top-level ML bindings are stored within the (global or local) theory - context. - - \item @{command "ML_prf"} is analogous to @{command "ML"} but works - within a proof context. Top-level ML bindings are stored within the - proof context in a purely sequential fashion, disregarding the - nested proof structure. ML bindings introduced by @{command - "ML_prf"} are discarded at the end of the proof. - - \item @{command "ML_val"} and @{command "ML_command"} are diagnostic - versions of @{command "ML"}, which means that the context may not be - updated. @{command "ML_val"} echos the bindings produced at the ML - toplevel, but @{command "ML_command"} is silent. - - \item @{command "setup"}~@{text "text"} changes the current theory - context by applying @{text "text"}, which refers to an ML expression - of type @{ML_type "theory -> theory"}. This enables to initialize - any object-logic specific tools and packages written in ML, for - example. - - \item @{command "local_setup"} is similar to @{command "setup"} for - a local theory context, and an ML expression of type @{ML_type - "local_theory -> local_theory"}. This allows to - invoke local theory specification packages without going through - concrete outer syntax, for example. - - \item @{command "attribute_setup"}~@{text "name = text description"} - defines an attribute in the current theory. The given @{text - "text"} has to be an ML expression of type - @{ML_type "attribute context_parser"}, cf.\ basic parsers defined in - structure @{ML_struct Args} and @{ML_struct Attrib}. - - In principle, attributes can operate both on a given theorem and the - implicit context, although in practice only one is modified and the - other serves as parameter. Here are examples for these two cases: - - \end{description} -*} - - attribute_setup my_rule = {* - Attrib.thms >> (fn ths => - Thm.rule_attribute - (fn context: Context.generic => fn th: thm => - let val th' = th OF ths - in th' end)) *} - - attribute_setup my_declaration = {* - Attrib.thms >> (fn ths => - Thm.declaration_attribute - (fn th: thm => fn context: Context.generic => - let val context' = context - in context' end)) *} - - -section {* Primitive specification elements *} - -subsection {* Type classes and sorts \label{sec:classes} *} - -text {* - \begin{matharray}{rcll} - @{command_def "classes"} & : & @{text "theory \ theory"} \\ - @{command_def "classrel"} & : & @{text "theory \ theory"} & (axiomatic!) \\ - @{command_def "default_sort"} & : & @{text "local_theory \ local_theory"} - \end{matharray} - - @{rail " - @@{command classes} (@{syntax classdecl} +) - ; - @@{command classrel} (@{syntax nameref} ('<' | '\') @{syntax nameref} + @'and') - ; - @@{command default_sort} @{syntax sort} - "} - - \begin{description} - - \item @{command "classes"}~@{text "c \ c\<^sub>1, \, c\<^sub>n"} declares class - @{text c} to be a subclass of existing classes @{text "c\<^sub>1, \, c\<^sub>n"}. - Isabelle implicitly maintains the transitive closure of the class - hierarchy. Cyclic class structures are not permitted. - - \item @{command "classrel"}~@{text "c\<^sub>1 \ c\<^sub>2"} states subclass - relations between existing classes @{text "c\<^sub>1"} and @{text "c\<^sub>2"}. - This is done axiomatically! The @{command_ref "subclass"} and - @{command_ref "instance"} commands (see \secref{sec:class}) provide - a way to introduce proven class relations. - - \item @{command "default_sort"}~@{text s} makes sort @{text s} the - new default sort for any type variable that is given explicitly in - the text, but lacks a sort constraint (wrt.\ the current context). - Type variables generated by type inference are not affected. - - Usually the default sort is only changed when defining a new - object-logic. For example, the default sort in Isabelle/HOL is - @{class type}, the class of all HOL types. - - When merging theories, the default sorts of the parents are - logically intersected, i.e.\ the representations as lists of classes - are joined. - - \end{description} -*} - - -subsection {* Types and type abbreviations \label{sec:types-pure} *} - -text {* - \begin{matharray}{rcll} - @{command_def "type_synonym"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "typedecl"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "arities"} & : & @{text "theory \ theory"} & (axiomatic!) \\ - \end{matharray} - - @{rail " - @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?) - ; - @@{command typedecl} @{syntax typespec} @{syntax mixfix}? - ; - @@{command arities} (@{syntax nameref} '::' @{syntax arity} +) - "} - - \begin{description} - - \item @{command "type_synonym"}~@{text "(\\<^sub>1, \, \\<^sub>n) t = \"} - introduces a \emph{type synonym} @{text "(\\<^sub>1, \, \\<^sub>n) t"} for the - existing type @{text "\"}. Unlike actual type definitions, as are - available in Isabelle/HOL for example, type synonyms are merely - syntactic abbreviations without any logical significance. - Internally, type synonyms are fully expanded. - - \item @{command "typedecl"}~@{text "(\\<^sub>1, \, \\<^sub>n) t"} declares a new - type constructor @{text t}. If the object-logic defines a base sort - @{text s}, then the constructor is declared to operate on that, via - the axiomatic specification @{command arities}~@{text "t :: (s, \, - s)s"}. - - \item @{command "arities"}~@{text "t :: (s\<^sub>1, \, s\<^sub>n)s"} augments - Isabelle's order-sorted signature of types by new type constructor - arities. This is done axiomatically! The @{command_ref "instantiation"} - target (see \secref{sec:class}) provides a way to introduce - proven type arities. - - \end{description} -*} - - -subsection {* Constants and definitions \label{sec:consts} *} - -text {* - \begin{matharray}{rcl} - @{command_def "consts"} & : & @{text "theory \ theory"} \\ - @{command_def "defs"} & : & @{text "theory \ theory"} \\ - \end{matharray} - - Definitions essentially express abbreviations within the logic. The - simplest form of a definition is @{text "c :: \ \ t"}, where @{text - c} is a newly declared constant. Isabelle also allows derived forms - where the arguments of @{text c} appear on the left, abbreviating a - prefix of @{text \}-abstractions, e.g.\ @{text "c \ \x y. t"} may be - written more conveniently as @{text "c x y \ t"}. Moreover, - definitions may be weakened by adding arbitrary pre-conditions: - @{text "A \ c x y \ t"}. - - \medskip The built-in well-formedness conditions for definitional - specifications are: - - \begin{itemize} - - \item Arguments (on the left-hand side) must be distinct variables. - - \item All variables on the right-hand side must also appear on the - left-hand side. - - \item All type variables on the right-hand side must also appear on - the left-hand side; this prohibits @{text "0 :: nat \ length ([] :: - \ list)"} for example. - - \item The definition must not be recursive. Most object-logics - provide definitional principles that can be used to express - recursion safely. - - \end{itemize} - - The right-hand side of overloaded definitions may mention overloaded constants - recursively at type instances corresponding to the immediate - argument types @{text "\\<^sub>1, \, \\<^sub>n"}. Incomplete - specification patterns impose global constraints on all occurrences, - e.g.\ @{text "d :: \ \ \"} on the left-hand side means that all - corresponding occurrences on some right-hand side need to be an - instance of this, general @{text "d :: \ \ \"} will be disallowed. - - @{rail " - @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +) - ; - @@{command defs} opt? (@{syntax axmdecl} @{syntax prop} +) - ; - opt: '(' @'unchecked'? @'overloaded'? ')' - "} - - \begin{description} - - \item @{command "consts"}~@{text "c :: \"} declares constant @{text - c} to have any instance of type scheme @{text \}. The optional - mixfix annotations may attach concrete syntax to the constants - declared. - - \item @{command "defs"}~@{text "name: eqn"} introduces @{text eqn} - as a definitional axiom for some existing constant. - - The @{text "(unchecked)"} option disables global dependency checks - for this definition, which is occasionally useful for exotic - overloading. It is at the discretion of the user to avoid malformed - theory specifications! - - The @{text "(overloaded)"} option declares definitions to be - potentially overloaded. Unless this option is given, a warning - message would be issued for any definitional equation with a more - special type than that of the corresponding constant declaration. - - \end{description} -*} - - -section {* Axioms and theorems \label{sec:axms-thms} *} - -text {* - \begin{matharray}{rcll} - @{command_def "axioms"} & : & @{text "theory \ theory"} & (axiomatic!) \\ - @{command_def "lemmas"} & : & @{text "local_theory \ local_theory"} \\ - @{command_def "theorems"} & : & @{text "local_theory \ local_theory"} \\ - \end{matharray} - - @{rail " - @@{command axioms} (@{syntax axmdecl} @{syntax prop} +) - ; - (@@{command lemmas} | @@{command theorems}) @{syntax target}? \\ - (@{syntax thmdef}? @{syntax thmrefs} + @'and') - (@'for' (@{syntax vars} + @'and'))? - "} - - \begin{description} - - \item @{command "axioms"}~@{text "a: \"} introduces arbitrary - statements as axioms of the meta-logic. In fact, axioms are - ``axiomatic theorems'', and may be referred later just as any other - theorem. - - Axioms are usually only introduced when declaring new logical - systems. Everyday work is typically done the hard way, with proper - definitions and proven theorems. - - \item @{command "lemmas"}~@{text "a = b\<^sub>1 \ b\<^sub>n"}~@{keyword_def - "for"}~@{text "x\<^sub>1 \ x\<^sub>m"} evaluates given facts (with attributes) in - the current context, which may be augmented by local variables. - Results are standardized before being stored, i.e.\ schematic - variables are renamed to enforce index @{text "0"} uniformly. - - \item @{command "theorems"} is the same as @{command "lemmas"}, but - marks the result as a different kind of facts. - - \end{description} -*} - - -section {* Oracles *} - -text {* - \begin{matharray}{rcll} - @{command_def "oracle"} & : & @{text "theory \ theory"} & (axiomatic!) \\ - \end{matharray} - - Oracles allow Isabelle to take advantage of external reasoners such - as arithmetic decision procedures, model checkers, fast tautology - checkers or computer algebra systems. Invoked as an oracle, an - external reasoner can create arbitrary Isabelle theorems. - - It is the responsibility of the user to ensure that the external - reasoner is as trustworthy as the application requires. Another - typical source of errors is the linkup between Isabelle and the - external tool, not just its concrete implementation, but also the - required translation between two different logical environments. - - Isabelle merely guarantees well-formedness of the propositions being - asserted, and records within the internal derivation object how - presumed theorems depend on unproven suppositions. - - @{rail " - @@{command oracle} @{syntax name} '=' @{syntax text} - "} - - \begin{description} - - \item @{command "oracle"}~@{text "name = text"} turns the given ML - expression @{text "text"} of type @{ML_text "'a -> cterm"} into an - ML function of type @{ML_text "'a -> thm"}, which is bound to the - global identifier @{ML_text name}. This acts like an infinitary - specification of axioms! Invoking the oracle only works within the - scope of the resulting theory. - - \end{description} - - See @{file "~~/src/HOL/ex/Iff_Oracle.thy"} for a worked example of - defining a new primitive rule as oracle, and turning it into a proof - method. -*} - - -section {* Name spaces *} - -text {* - \begin{matharray}{rcl} - @{command_def "hide_class"} & : & @{text "theory \ theory"} \\ - @{command_def "hide_type"} & : & @{text "theory \ theory"} \\ - @{command_def "hide_const"} & : & @{text "theory \ theory"} \\ - @{command_def "hide_fact"} & : & @{text "theory \ theory"} \\ - \end{matharray} - - @{rail " - ( @{command hide_class} | @{command hide_type} | - @{command hide_const} | @{command hide_fact} ) ('(' @'open' ')')? (@{syntax nameref} + ) - "} - - Isabelle organizes any kind of name declarations (of types, - constants, theorems etc.) by separate hierarchically structured name - spaces. Normally the user does not have to control the behavior of - name spaces by hand, yet the following commands provide some way to - do so. - - \begin{description} - - \item @{command "hide_class"}~@{text names} fully removes class - declarations from a given name space; with the @{text "(open)"} - option, only the base name is hidden. - - Note that hiding name space accesses has no impact on logical - declarations --- they remain valid internally. Entities that are no - longer accessible to the user are printed with the special qualifier - ``@{text "??"}'' prefixed to the full internal name. - - \item @{command "hide_type"}, @{command "hide_const"}, and @{command - "hide_fact"} are similar to @{command "hide_class"}, but hide types, - constants, and facts, respectively. - - \end{description} -*} - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Symbols.thy --- a/doc-src/IsarRef/Thy/Symbols.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,46 +0,0 @@ -theory Symbols -imports Base Main -begin - -chapter {* Predefined Isabelle symbols \label{app:symbols} *} - -text {* - Isabelle supports an infinite number of non-ASCII symbols, which are - represented in source text as @{verbatim "\\"}@{verbatim "<"}@{text - name}@{verbatim ">"} (where @{text name} may be any identifier). It - is left to front-end tools how to present these symbols to the user. - The collection of predefined standard symbols given below is - available by default for Isabelle document output, due to - appropriate definitions of @{verbatim "\\"}@{verbatim isasym}@{text - name} for each @{verbatim "\\"}@{verbatim "<"}@{text name}@{verbatim - ">"} in the @{verbatim isabellesym.sty} file. Most of these symbols - are displayed properly in Proof~General and Isabelle/jEdit. - - Moreover, any single symbol (or ASCII character) may be prefixed by - @{verbatim "\\"}@{verbatim "<^sup>"}, for superscript and @{verbatim - "\\"}@{verbatim "<^sub>"}, for subscript, such as @{verbatim - "A\\"}@{verbatim "<^sup>\"}, for @{text "A\<^sup>\"} the alternative - versions @{verbatim "\\"}@{verbatim "<^isub>"} and @{verbatim - "\\"}@{verbatim "<^isup>"} are considered as quasi letters and may - be used within identifiers. Sub- and superscripts that span a - region of text are marked up with @{verbatim "\\"}@{verbatim - "<^bsub>"}@{text "\"}@{verbatim "\\"}@{verbatim "<^esub>"}, and - @{verbatim "\\"}@{verbatim "<^bsup>"}@{text "\"}@{verbatim - "\\"}@{verbatim "<^esup>"} respectively. Furthermore, all ASCII - characters and most other symbols may be printed in bold by - prefixing @{verbatim "\\"}@{verbatim "<^bold>"} such as @{verbatim - "\\"}@{verbatim "<^bold>\\"}@{verbatim ""} for @{text - "\<^bold>\"}. Note that @{verbatim "\\"}@{verbatim "<^bold>"}, may - \emph{not} be combined with sub- or superscripts for single symbols. - - Further details of Isabelle document preparation are covered in - \chref{ch:document-prep}. - - \begin{center} - \begin{isabellebody} - \input{syms} - \end{isabellebody} - \end{center} -*} - -end \ No newline at end of file diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/Synopsis.thy --- a/doc-src/IsarRef/Thy/Synopsis.thy Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1108 +0,0 @@ -theory Synopsis -imports Base Main -begin - -chapter {* Synopsis *} - -section {* Notepad *} - -text {* - An Isar proof body serves as mathematical notepad to compose logical - content, consisting of types, terms, facts. -*} - - -subsection {* Types and terms *} - -notepad -begin - txt {* Locally fixed entities: *} - fix x -- {* local constant, without any type information yet *} - fix x :: 'a -- {* variant with explicit type-constraint for subsequent use*} - - fix a b - assume "a = b" -- {* type assignment at first occurrence in concrete term *} - - txt {* Definitions (non-polymorphic): *} - def x \ "t::'a" - - txt {* Abbreviations (polymorphic): *} - let ?f = "\x. x" - term "?f ?f" - - txt {* Notation: *} - write x ("***") -end - - -subsection {* Facts *} - -text {* - A fact is a simultaneous list of theorems. -*} - - -subsubsection {* Producing facts *} - -notepad -begin - - txt {* Via assumption (``lambda''): *} - assume a: A - - txt {* Via proof (``let''): *} - have b: B sorry - - txt {* Via abbreviation (``let''): *} - note c = a b - -end - - -subsubsection {* Referencing facts *} - -notepad -begin - txt {* Via explicit name: *} - assume a: A - note a - - txt {* Via implicit name: *} - assume A - note this - - txt {* Via literal proposition (unification with results from the proof text): *} - assume A - note `A` - - assume "\x. B x" - note `B a` - note `B b` -end - - -subsubsection {* Manipulating facts *} - -notepad -begin - txt {* Instantiation: *} - assume a: "\x. B x" - note a - note a [of b] - note a [where x = b] - - txt {* Backchaining: *} - assume 1: A - assume 2: "A \ C" - note 2 [OF 1] - note 1 [THEN 2] - - txt {* Symmetric results: *} - assume "x = y" - note this [symmetric] - - assume "x \ y" - note this [symmetric] - - txt {* Adhoc-simplification (take care!): *} - assume "P ([] @ xs)" - note this [simplified] -end - - -subsubsection {* Projections *} - -text {* - Isar facts consist of multiple theorems. There is notation to project - interval ranges. -*} - -notepad -begin - assume stuff: A B C D - note stuff(1) - note stuff(2-3) - note stuff(2-) -end - - -subsubsection {* Naming conventions *} - -text {* - \begin{itemize} - - \item Lower-case identifiers are usually preferred. - - \item Facts can be named after the main term within the proposition. - - \item Facts should \emph{not} be named after the command that - introduced them (@{command "assume"}, @{command "have"}). This is - misleading and hard to maintain. - - \item Natural numbers can be used as ``meaningless'' names (more - appropriate than @{text "a1"}, @{text "a2"} etc.) - - \item Symbolic identifiers are supported (e.g. @{text "*"}, @{text - "**"}, @{text "***"}). - - \end{itemize} -*} - - -subsection {* Block structure *} - -text {* - The formal notepad is block structured. The fact produced by the last - entry of a block is exported into the outer context. -*} - -notepad -begin - { - have a: A sorry - have b: B sorry - note a b - } - note this - note `A` - note `B` -end - -text {* Explicit blocks as well as implicit blocks of nested goal - statements (e.g.\ @{command have}) automatically introduce one extra - pair of parentheses in reserve. The @{command next} command allows - to ``jump'' between these sub-blocks. *} - -notepad -begin - - { - have a: A sorry - next - have b: B - proof - - show B sorry - next - have c: C sorry - next - have d: D sorry - qed - } - - txt {* Alternative version with explicit parentheses everywhere: *} - - { - { - have a: A sorry - } - { - have b: B - proof - - { - show B sorry - } - { - have c: C sorry - } - { - have d: D sorry - } - qed - } - } - -end - - -section {* Calculational reasoning \label{sec:calculations-synopsis} *} - -text {* - For example, see @{file "~~/src/HOL/Isar_Examples/Group.thy"}. -*} - - -subsection {* Special names in Isar proofs *} - -text {* - \begin{itemize} - - \item term @{text "?thesis"} --- the main conclusion of the - innermost pending claim - - \item term @{text "\"} --- the argument of the last explicitly - stated result (for infix application this is the right-hand side) - - \item fact @{text "this"} --- the last result produced in the text - - \end{itemize} -*} - -notepad -begin - have "x = y" - proof - - term ?thesis - show ?thesis sorry - term ?thesis -- {* static! *} - qed - term "\" - thm this -end - -text {* Calculational reasoning maintains the special fact called - ``@{text calculation}'' in the background. Certain language - elements combine primary @{text this} with secondary @{text - calculation}. *} - - -subsection {* Transitive chains *} - -text {* The Idea is to combine @{text this} and @{text calculation} - via typical @{text trans} rules (see also @{command - print_trans_rules}): *} - -thm trans -thm less_trans -thm less_le_trans - -notepad -begin - txt {* Plain bottom-up calculation: *} - have "a = b" sorry - also - have "b = c" sorry - also - have "c = d" sorry - finally - have "a = d" . - - txt {* Variant using the @{text "\"} abbreviation: *} - have "a = b" sorry - also - have "\ = c" sorry - also - have "\ = d" sorry - finally - have "a = d" . - - txt {* Top-down version with explicit claim at the head: *} - have "a = d" - proof - - have "a = b" sorry - also - have "\ = c" sorry - also - have "\ = d" sorry - finally - show ?thesis . - qed -next - txt {* Mixed inequalities (require suitable base type): *} - fix a b c d :: nat - - have "a < b" sorry - also - have "b \ c" sorry - also - have "c = d" sorry - finally - have "a < d" . -end - - -subsubsection {* Notes *} - -text {* - \begin{itemize} - - \item The notion of @{text trans} rule is very general due to the - flexibility of Isabelle/Pure rule composition. - - \item User applications may declare their own rules, with some care - about the operational details of higher-order unification. - - \end{itemize} -*} - - -subsection {* Degenerate calculations and bigstep reasoning *} - -text {* The Idea is to append @{text this} to @{text calculation}, - without rule composition. *} - -notepad -begin - txt {* A vacuous proof: *} - have A sorry - moreover - have B sorry - moreover - have C sorry - ultimately - have A and B and C . -next - txt {* Slightly more content (trivial bigstep reasoning): *} - have A sorry - moreover - have B sorry - moreover - have C sorry - ultimately - have "A \ B \ C" by blast -next - txt {* More ambitious bigstep reasoning involving structured results: *} - have "A \ B \ C" sorry - moreover - { assume A have R sorry } - moreover - { assume B have R sorry } - moreover - { assume C have R sorry } - ultimately - have R by blast -- {* ``big-bang integration'' of proof blocks (occasionally fragile) *} -end - - -section {* Induction *} - -subsection {* Induction as Natural Deduction *} - -text {* In principle, induction is just a special case of Natural - Deduction (see also \secref{sec:natural-deduction-synopsis}). For - example: *} - -thm nat.induct -print_statement nat.induct - -notepad -begin - fix n :: nat - have "P n" - proof (rule nat.induct) -- {* fragile rule application! *} - show "P 0" sorry - next - fix n :: nat - assume "P n" - show "P (Suc n)" sorry - qed -end - -text {* - In practice, much more proof infrastructure is required. - - The proof method @{method induct} provides: - \begin{itemize} - - \item implicit rule selection and robust instantiation - - \item context elements via symbolic case names - - \item support for rule-structured induction statements, with local - parameters, premises, etc. - - \end{itemize} -*} - -notepad -begin - fix n :: nat - have "P n" - proof (induct n) - case 0 - show ?case sorry - next - case (Suc n) - from Suc.hyps show ?case sorry - qed -end - - -subsubsection {* Example *} - -text {* - The subsequent example combines the following proof patterns: - \begin{itemize} - - \item outermost induction (over the datatype structure of natural - numbers), to decompose the proof problem in top-down manner - - \item calculational reasoning (\secref{sec:calculations-synopsis}) - to compose the result in each case - - \item solving local claims within the calculation by simplification - - \end{itemize} -*} - -lemma - fixes n :: nat - shows "(\i=0..n. i) = n * (n + 1) div 2" -proof (induct n) - case 0 - have "(\i=0..0. i) = (0::nat)" by simp - also have "\ = 0 * (0 + 1) div 2" by simp - finally show ?case . -next - case (Suc n) - have "(\i=0..Suc n. i) = (\i=0..n. i) + (n + 1)" by simp - also have "\ = n * (n + 1) div 2 + (n + 1)" by (simp add: Suc.hyps) - also have "\ = (n * (n + 1) + 2 * (n + 1)) div 2" by simp - also have "\ = (Suc n * (Suc n + 1)) div 2" by simp - finally show ?case . -qed - -text {* This demonstrates how induction proofs can be done without - having to consider the raw Natural Deduction structure. *} - - -subsection {* Induction with local parameters and premises *} - -text {* Idea: Pure rule statements are passed through the induction - rule. This achieves convenient proof patterns, thanks to some - internal trickery in the @{method induct} method. - - Important: Using compact HOL formulae with @{text "\/\"} is a - well-known anti-pattern! It would produce useless formal noise. -*} - -notepad -begin - fix n :: nat - fix P :: "nat \ bool" - fix Q :: "'a \ nat \ bool" - - have "P n" - proof (induct n) - case 0 - show "P 0" sorry - next - case (Suc n) - from `P n` show "P (Suc n)" sorry - qed - - have "A n \ P n" - proof (induct n) - case 0 - from `A 0` show "P 0" sorry - next - case (Suc n) - from `A n \ P n` - and `A (Suc n)` show "P (Suc n)" sorry - qed - - have "\x. Q x n" - proof (induct n) - case 0 - show "Q x 0" sorry - next - case (Suc n) - from `\x. Q x n` show "Q x (Suc n)" sorry - txt {* Local quantification admits arbitrary instances: *} - note `Q a n` and `Q b n` - qed -end - - -subsection {* Implicit induction context *} - -text {* The @{method induct} method can isolate local parameters and - premises directly from the given statement. This is convenient in - practical applications, but requires some understanding of what is - going on internally (as explained above). *} - -notepad -begin - fix n :: nat - fix Q :: "'a \ nat \ bool" - - fix x :: 'a - assume "A x n" - then have "Q x n" - proof (induct n arbitrary: x) - case 0 - from `A x 0` show "Q x 0" sorry - next - case (Suc n) - from `\x. A x n \ Q x n` -- {* arbitrary instances can be produced here *} - and `A x (Suc n)` show "Q x (Suc n)" sorry - qed -end - - -subsection {* Advanced induction with term definitions *} - -text {* Induction over subexpressions of a certain shape are delicate - to formalize. The Isar @{method induct} method provides - infrastructure for this. - - Idea: sub-expressions of the problem are turned into a defined - induction variable; often accompanied with fixing of auxiliary - parameters in the original expression. *} - -notepad -begin - fix a :: "'a \ nat" - fix A :: "nat \ bool" - - assume "A (a x)" - then have "P (a x)" - proof (induct "a x" arbitrary: x) - case 0 - note prem = `A (a x)` - and defn = `0 = a x` - show "P (a x)" sorry - next - case (Suc n) - note hyp = `\x. n = a x \ A (a x) \ P (a x)` - and prem = `A (a x)` - and defn = `Suc n = a x` - show "P (a x)" sorry - qed -end - - -section {* Natural Deduction \label{sec:natural-deduction-synopsis} *} - -subsection {* Rule statements *} - -text {* - Isabelle/Pure ``theorems'' are always natural deduction rules, - which sometimes happen to consist of a conclusion only. - - The framework connectives @{text "\"} and @{text "\"} indicate the - rule structure declaratively. For example: *} - -thm conjI -thm impI -thm nat.induct - -text {* - The object-logic is embedded into the Pure framework via an implicit - derivability judgment @{term "Trueprop :: bool \ prop"}. - - Thus any HOL formulae appears atomic to the Pure framework, while - the rule structure outlines the corresponding proof pattern. - - This can be made explicit as follows: -*} - -notepad -begin - write Trueprop ("Tr") - - thm conjI - thm impI - thm nat.induct -end - -text {* - Isar provides first-class notation for rule statements as follows. -*} - -print_statement conjI -print_statement impI -print_statement nat.induct - - -subsubsection {* Examples *} - -text {* - Introductions and eliminations of some standard connectives of - the object-logic can be written as rule statements as follows. (The - proof ``@{command "by"}~@{method blast}'' serves as sanity check.) -*} - -lemma "(P \ False) \ \ P" by blast -lemma "\ P \ P \ Q" by blast - -lemma "P \ Q \ P \ Q" by blast -lemma "P \ Q \ (P \ Q \ R) \ R" by blast - -lemma "P \ P \ Q" by blast -lemma "Q \ P \ Q" by blast -lemma "P \ Q \ (P \ R) \ (Q \ R) \ R" by blast - -lemma "(\x. P x) \ (\x. P x)" by blast -lemma "(\x. P x) \ P x" by blast - -lemma "P x \ (\x. P x)" by blast -lemma "(\x. P x) \ (\x. P x \ R) \ R" by blast - -lemma "x \ A \ x \ B \ x \ A \ B" by blast -lemma "x \ A \ B \ (x \ A \ x \ B \ R) \ R" by blast - -lemma "x \ A \ x \ A \ B" by blast -lemma "x \ B \ x \ A \ B" by blast -lemma "x \ A \ B \ (x \ A \ R) \ (x \ B \ R) \ R" by blast - - -subsection {* Isar context elements *} - -text {* We derive some results out of the blue, using Isar context - elements and some explicit blocks. This illustrates their meaning - wrt.\ Pure connectives, without goal states getting in the way. *} - -notepad -begin - { - fix x - have "B x" sorry - } - have "\x. B x" by fact - -next - - { - assume A - have B sorry - } - have "A \ B" by fact - -next - - { - def x \ t - have "B x" sorry - } - have "B t" by fact - -next - - { - obtain x :: 'a where "B x" sorry - have C sorry - } - have C by fact - -end - - -subsection {* Pure rule composition *} - -text {* - The Pure framework provides means for: - - \begin{itemize} - - \item backward-chaining of rules by @{inference resolution} - - \item closing of branches by @{inference assumption} - - \end{itemize} - - Both principles involve higher-order unification of @{text \}-terms - modulo @{text "\\\"}-equivalence (cf.\ Huet and Miller). *} - -notepad -begin - assume a: A and b: B - thm conjI - thm conjI [of A B] -- "instantiation" - thm conjI [of A B, OF a b] -- "instantiation and composition" - thm conjI [OF a b] -- "composition via unification (trivial)" - thm conjI [OF `A` `B`] - - thm conjI [OF disjI1] -end - -text {* Note: Low-level rule composition is tedious and leads to - unreadable~/ unmaintainable expressions in the text. *} - - -subsection {* Structured backward reasoning *} - -text {* Idea: Canonical proof decomposition via @{command fix}~/ - @{command assume}~/ @{command show}, where the body produces a - natural deduction rule to refine some goal. *} - -notepad -begin - fix A B :: "'a \ bool" - - have "\x. A x \ B x" - proof - - fix x - assume "A x" - show "B x" sorry - qed - - have "\x. A x \ B x" - proof - - { - fix x - assume "A x" - show "B x" sorry - } -- "implicit block structure made explicit" - note `\x. A x \ B x` - -- "side exit for the resulting rule" - qed -end - - -subsection {* Structured rule application *} - -text {* - Idea: Previous facts and new claims are composed with a rule from - the context (or background library). -*} - -notepad -begin - assume r1: "A \ B \ C" -- {* simple rule (Horn clause) *} - - have A sorry -- "prefix of facts via outer sub-proof" - then have C - proof (rule r1) - show B sorry -- "remaining rule premises via inner sub-proof" - qed - - have C - proof (rule r1) - show A sorry - show B sorry - qed - - have A and B sorry - then have C - proof (rule r1) - qed - - have A and B sorry - then have C by (rule r1) - -next - - assume r2: "A \ (\x. B1 x \ B2 x) \ C" -- {* nested rule *} - - have A sorry - then have C - proof (rule r2) - fix x - assume "B1 x" - show "B2 x" sorry - qed - - txt {* The compound rule premise @{prop "\x. B1 x \ B2 x"} is better - addressed via @{command fix}~/ @{command assume}~/ @{command show} - in the nested proof body. *} -end - - -subsection {* Example: predicate logic *} - -text {* - Using the above principles, standard introduction and elimination proofs - of predicate logic connectives of HOL work as follows. -*} - -notepad -begin - have "A \ B" and A sorry - then have B .. - - have A sorry - then have "A \ B" .. - - have B sorry - then have "A \ B" .. - - have "A \ B" sorry - then have C - proof - assume A - then show C sorry - next - assume B - then show C sorry - qed - - have A and B sorry - then have "A \ B" .. - - have "A \ B" sorry - then have A .. - - have "A \ B" sorry - then have B .. - - have False sorry - then have A .. - - have True .. - - have "\ A" - proof - assume A - then show False sorry - qed - - have "\ A" and A sorry - then have B .. - - have "\x. P x" - proof - fix x - show "P x" sorry - qed - - have "\x. P x" sorry - then have "P a" .. - - have "\x. P x" - proof - show "P a" sorry - qed - - have "\x. P x" sorry - then have C - proof - fix a - assume "P a" - show C sorry - qed - - txt {* Less awkward version using @{command obtain}: *} - have "\x. P x" sorry - then obtain a where "P a" .. -end - -text {* Further variations to illustrate Isar sub-proofs involving - @{command show}: *} - -notepad -begin - have "A \ B" - proof -- {* two strictly isolated subproofs *} - show A sorry - next - show B sorry - qed - - have "A \ B" - proof -- {* one simultaneous sub-proof *} - show A and B sorry - qed - - have "A \ B" - proof -- {* two subproofs in the same context *} - show A sorry - show B sorry - qed - - have "A \ B" - proof -- {* swapped order *} - show B sorry - show A sorry - qed - - have "A \ B" - proof -- {* sequential subproofs *} - show A sorry - show B using `A` sorry - qed -end - - -subsubsection {* Example: set-theoretic operators *} - -text {* There is nothing special about logical connectives (@{text - "\"}, @{text "\"}, @{text "\"}, @{text "\"} etc.). Operators from - set-theory or lattice-theory work analogously. It is only a matter - of rule declarations in the library; rules can be also specified - explicitly. -*} - -notepad -begin - have "x \ A" and "x \ B" sorry - then have "x \ A \ B" .. - - have "x \ A" sorry - then have "x \ A \ B" .. - - have "x \ B" sorry - then have "x \ A \ B" .. - - have "x \ A \ B" sorry - then have C - proof - assume "x \ A" - then show C sorry - next - assume "x \ B" - then show C sorry - qed - -next - have "x \ \A" - proof - fix a - assume "a \ A" - show "x \ a" sorry - qed - - have "x \ \A" sorry - then have "x \ a" - proof - show "a \ A" sorry - qed - - have "a \ A" and "x \ a" sorry - then have "x \ \A" .. - - have "x \ \A" sorry - then obtain a where "a \ A" and "x \ a" .. -end - - -section {* Generalized elimination and cases *} - -subsection {* General elimination rules *} - -text {* - The general format of elimination rules is illustrated by the - following typical representatives: -*} - -thm exE -- {* local parameter *} -thm conjE -- {* local premises *} -thm disjE -- {* split into cases *} - -text {* - Combining these characteristics leads to the following general scheme - for elimination rules with cases: - - \begin{itemize} - - \item prefix of assumptions (or ``major premises'') - - \item one or more cases that enable to establish the main conclusion - in an augmented context - - \end{itemize} -*} - -notepad -begin - assume r: - "A1 \ A2 \ (* assumptions *) - (\x y. B1 x y \ C1 x y \ R) \ (* case 1 *) - (\x y. B2 x y \ C2 x y \ R) \ (* case 2 *) - R (* main conclusion *)" - - have A1 and A2 sorry - then have R - proof (rule r) - fix x y - assume "B1 x y" and "C1 x y" - show ?thesis sorry - next - fix x y - assume "B2 x y" and "C2 x y" - show ?thesis sorry - qed -end - -text {* Here @{text "?thesis"} is used to refer to the unchanged goal - statement. *} - - -subsection {* Rules with cases *} - -text {* - Applying an elimination rule to some goal, leaves that unchanged - but allows to augment the context in the sub-proof of each case. - - Isar provides some infrastructure to support this: - - \begin{itemize} - - \item native language elements to state eliminations - - \item symbolic case names - - \item method @{method cases} to recover this structure in a - sub-proof - - \end{itemize} -*} - -print_statement exE -print_statement conjE -print_statement disjE - -lemma - assumes A1 and A2 -- {* assumptions *} - obtains - (case1) x y where "B1 x y" and "C1 x y" - | (case2) x y where "B2 x y" and "C2 x y" - sorry - - -subsubsection {* Example *} - -lemma tertium_non_datur: - obtains - (T) A - | (F) "\ A" - by blast - -notepad -begin - fix x y :: 'a - have C - proof (cases "x = y" rule: tertium_non_datur) - case T - from `x = y` show ?thesis sorry - next - case F - from `x \ y` show ?thesis sorry - qed -end - - -subsubsection {* Example *} - -text {* - Isabelle/HOL specification mechanisms (datatype, inductive, etc.) - provide suitable derived cases rules. -*} - -datatype foo = Foo | Bar foo - -notepad -begin - fix x :: foo - have C - proof (cases x) - case Foo - from `x = Foo` show ?thesis sorry - next - case (Bar a) - from `x = Bar a` show ?thesis sorry - qed -end - - -subsection {* Obtaining local contexts *} - -text {* A single ``case'' branch may be inlined into Isar proof text - via @{command obtain}. This proves @{prop "(\x. B x \ thesis) \ - thesis"} on the spot, and augments the context afterwards. *} - -notepad -begin - fix B :: "'a \ bool" - - obtain x where "B x" sorry - note `B x` - - txt {* Conclusions from this context may not mention @{term x} again! *} - { - obtain x where "B x" sorry - from `B x` have C sorry - } - note `C` -end - -end diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Document_Preparation.tex --- a/doc-src/IsarRef/Thy/document/Document_Preparation.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,990 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Document{\isaliteral{5F}{\isacharunderscore}}Preparation}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Document{\isaliteral{5F}{\isacharunderscore}}Preparation\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Document preparation \label{ch:document-prep}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/Isar provides a simple document preparation system - based on regular {PDF-\LaTeX} technology, with full support for - hyper-links and bookmarks. Thus the results are well suited for WWW - browsing and as printed copies. - - \medskip Isabelle generates {\LaTeX} output while running a - \emph{logic session} (see also \cite{isabelle-sys}). Getting - started with a working configuration for common situations is quite - easy by using the Isabelle \verb|mkdir| and \verb|make| - tools. First invoke -\begin{ttbox} - isabelle mkdir Foo -\end{ttbox} - to initialize a separate directory for session \verb|Foo| (it - is safe to experiment, since \verb|isabelle mkdir| never - overwrites existing files). Ensure that \verb|Foo/ROOT.ML| - holds ML commands to load all theories required for this session; - furthermore \verb|Foo/document/root.tex| should include any - special {\LaTeX} macro packages required for your document (the - default is usually sufficient as a start). - - The session is controlled by a separate \verb|IsaMakefile| - (with crude source dependencies by default). This file is located - one level up from the \verb|Foo| directory location. Now - invoke -\begin{ttbox} - isabelle make Foo -\end{ttbox} - to run the \verb|Foo| session, with browser information and - document preparation enabled. Unless any errors are reported by - Isabelle or {\LaTeX}, the output will appear inside the directory - defined by the \verb|ISABELLE_BROWSER_INFO| setting (as - reported by the batch job in verbose mode). - - \medskip You may also consider to tune the \verb|usedir| - options in \verb|IsaMakefile|, for example to switch the output - format between \verb|pdf| and \verb|dvi|, or activate the - \verb|-D| option to retain a second copy of the generated - {\LaTeX} sources (for manual inspection or separate runs of - \hyperlink{executable.latex}{\mbox{\isa{\isatt{latex}}}}). - - \medskip See \emph{The Isabelle System Manual} \cite{isabelle-sys} - for further details on Isabelle logic sessions and theory - presentation. The Isabelle/HOL tutorial \cite{isabelle-hol-book} - also covers theory presentation to some extent.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Markup commands \label{sec:markup}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{header}\hypertarget{command.header}{\hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}toplevel\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ toplevel{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \indexdef{}{command}{chapter}\hypertarget{command.chapter}{\hyperlink{command.chapter}{\mbox{\isa{\isacommand{chapter}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{section}\hypertarget{command.section}{\hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{subsection}\hypertarget{command.subsection}{\hyperlink{command.subsection}{\mbox{\isa{\isacommand{subsection}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{subsubsection}\hypertarget{command.subsubsection}{\hyperlink{command.subsubsection}{\mbox{\isa{\isacommand{subsubsection}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{text}\hypertarget{command.text}{\hyperlink{command.text}{\mbox{\isa{\isacommand{text}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{text\_raw}\hypertarget{command.text-raw}{\hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isaliteral{5F}{\isacharunderscore}}raw}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \indexdef{}{command}{sect}\hypertarget{command.sect}{\hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{subsect}\hypertarget{command.subsect}{\hyperlink{command.subsect}{\mbox{\isa{\isacommand{subsect}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{subsubsect}\hypertarget{command.subsubsect}{\hyperlink{command.subsubsect}{\mbox{\isa{\isacommand{subsubsect}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{txt}\hypertarget{command.txt}{\hyperlink{command.txt}{\mbox{\isa{\isacommand{txt}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{txt\_raw}\hypertarget{command.txt-raw}{\hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isaliteral{5F}{\isacharunderscore}}raw}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Markup commands provide a structured way to insert text into the - document generated from a theory. Each markup command takes a - single \hyperlink{syntax.text}{\mbox{\isa{text}}} argument, which is passed as argument to a - corresponding {\LaTeX} macro. The default macros provided by - \verb|~~/lib/texinputs/isabelle.sty| can be redefined according - to the needs of the underlying document and {\LaTeX} styles. - - Note that formal comments (\secref{sec:comments}) are similar to - markup commands, but have a different status within Isabelle/Isar - syntax. - - \begin{railoutput} -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{command.chapter}{\mbox{\isa{\isacommand{chapter}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.subsection}{\mbox{\isa{\isacommand{subsection}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.subsubsection}{\mbox{\isa{\isacommand{subsubsection}}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{command.text}{\mbox{\isa{\isacommand{text}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\rail@begin{7}{} -\rail@bar -\rail@term{\hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isaliteral{5F}{\isacharunderscore}}raw}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.subsect}{\mbox{\isa{\isacommand{subsect}}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{command.subsubsect}{\mbox{\isa{\isacommand{subsubsect}}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{command.txt}{\mbox{\isa{\isacommand{txt}}}}}[] -\rail@nextbar{6} -\rail@term{\hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isaliteral{5F}{\isacharunderscore}}raw}}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}} provides plain text markup just preceding - the formal beginning of a theory. The corresponding {\LaTeX} macro - is \verb|\isamarkupheader|, which acts like \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}} by default. - - \item \hyperlink{command.chapter}{\mbox{\isa{\isacommand{chapter}}}}, \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}}, \hyperlink{command.subsection}{\mbox{\isa{\isacommand{subsection}}}}, - and \hyperlink{command.subsubsection}{\mbox{\isa{\isacommand{subsubsection}}}} mark chapter and section headings - within the main theory body or local theory targets. The - corresponding {\LaTeX} macros are \verb|\isamarkupchapter|, - \verb|\isamarkupsection|, \verb|\isamarkupsubsection| etc. - - \item \hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}, \hyperlink{command.subsect}{\mbox{\isa{\isacommand{subsect}}}}, and \hyperlink{command.subsubsect}{\mbox{\isa{\isacommand{subsubsect}}}} - mark section headings within proofs. The corresponding {\LaTeX} - macros are \verb|\isamarkupsect|, \verb|\isamarkupsubsect| etc. - - \item \hyperlink{command.text}{\mbox{\isa{\isacommand{text}}}} and \hyperlink{command.txt}{\mbox{\isa{\isacommand{txt}}}} specify paragraphs of plain - text. This corresponds to a {\LaTeX} environment \verb|\begin{isamarkuptext}| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|\end{isamarkuptext}| etc. - - \item \hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isaliteral{5F}{\isacharunderscore}}raw}}}} and \hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isaliteral{5F}{\isacharunderscore}}raw}}}} insert {\LaTeX} - source into the output, without additional markup. Thus the full - range of document manipulations becomes available, at the risk of - messing up document output. - - \end{description} - - Except for \hyperlink{command.text-raw}{\mbox{\isa{\isacommand{text{\isaliteral{5F}{\isacharunderscore}}raw}}}} and \hyperlink{command.txt-raw}{\mbox{\isa{\isacommand{txt{\isaliteral{5F}{\isacharunderscore}}raw}}}}, the text - passed to any of the above markup commands may refer to formal - entities via \emph{document antiquotations}, see also - \secref{sec:antiq}. These are interpreted in the present theory or - proof context, or the named \isa{{\isaliteral{22}{\isachardoublequote}}target{\isaliteral{22}{\isachardoublequote}}}. - - \medskip The proof markup commands closely resemble those for theory - specifications, but have a different formal status and produce - different {\LaTeX} macros. The default definitions coincide for - analogous commands such as \hyperlink{command.section}{\mbox{\isa{\isacommand{section}}}} and \hyperlink{command.sect}{\mbox{\isa{\isacommand{sect}}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Document Antiquotations \label{sec:antiq}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{antiquotation}{theory}\hypertarget{antiquotation.theory}{\hyperlink{antiquotation.theory}{\mbox{\isa{theory}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{thm}\hypertarget{antiquotation.thm}{\hyperlink{antiquotation.thm}{\mbox{\isa{thm}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{lemma}\hypertarget{antiquotation.lemma}{\hyperlink{antiquotation.lemma}{\mbox{\isa{lemma}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{prop}\hypertarget{antiquotation.prop}{\hyperlink{antiquotation.prop}{\mbox{\isa{prop}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{term}\hypertarget{antiquotation.term}{\hyperlink{antiquotation.term}{\mbox{\isa{term}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{term\_type}\hypertarget{antiquotation.term-type}{\hyperlink{antiquotation.term-type}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}type}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{typeof}\hypertarget{antiquotation.typeof}{\hyperlink{antiquotation.typeof}{\mbox{\isa{typeof}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{const}\hypertarget{antiquotation.const}{\hyperlink{antiquotation.const}{\mbox{\isa{const}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{abbrev}\hypertarget{antiquotation.abbrev}{\hyperlink{antiquotation.abbrev}{\mbox{\isa{abbrev}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{typ}\hypertarget{antiquotation.typ}{\hyperlink{antiquotation.typ}{\mbox{\isa{typ}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{type}\hypertarget{antiquotation.type}{\hyperlink{antiquotation.type}{\mbox{\isa{type}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{class}\hypertarget{antiquotation.class}{\hyperlink{antiquotation.class}{\mbox{\isa{class}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{text}\hypertarget{antiquotation.text}{\hyperlink{antiquotation.text}{\mbox{\isa{text}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{goals}\hypertarget{antiquotation.goals}{\hyperlink{antiquotation.goals}{\mbox{\isa{goals}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{subgoals}\hypertarget{antiquotation.subgoals}{\hyperlink{antiquotation.subgoals}{\mbox{\isa{subgoals}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{prf}\hypertarget{antiquotation.prf}{\hyperlink{antiquotation.prf}{\mbox{\isa{prf}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{full\_prf}\hypertarget{antiquotation.full-prf}{\hyperlink{antiquotation.full-prf}{\mbox{\isa{full{\isaliteral{5F}{\isacharunderscore}}prf}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{ML}\hypertarget{antiquotation.ML}{\hyperlink{antiquotation.ML}{\mbox{\isa{ML}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{ML\_op}\hypertarget{antiquotation.ML-op}{\hyperlink{antiquotation.ML-op}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}op}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{ML\_type}\hypertarget{antiquotation.ML-type}{\hyperlink{antiquotation.ML-type}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}type}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{ML\_struct}\hypertarget{antiquotation.ML-struct}{\hyperlink{antiquotation.ML-struct}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}struct}}}} & : & \isa{antiquotation} \\ - \indexdef{}{antiquotation}{file}\hypertarget{antiquotation.file}{\hyperlink{antiquotation.file}{\mbox{\isa{file}}}} & : & \isa{antiquotation} \\ - \end{matharray} - - The overall content of an Isabelle/Isar theory may alternate between - formal and informal text. The main body consists of formal - specification and proof commands, interspersed with markup commands - (\secref{sec:markup}) or document comments (\secref{sec:comments}). - The argument of markup commands quotes informal text to be printed - in the resulting document, but may again refer to formal entities - via \emph{document antiquotations}. - - For example, embedding of ``\isa{{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}term\ {\isaliteral{5B}{\isacharbrackleft}}show{\isaliteral{5F}{\isacharunderscore}}types{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{22}{\isachardoublequote}}f\ x\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{2B}{\isacharplus}}\ x{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7D}{\isacharbraceright}}}'' - within a text block makes - \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document. - - Antiquotations usually spare the author tedious typing of logical - entities in full detail. Even more importantly, some degree of - consistency-checking between the main body of formal text and its - informal explanation is achieved, since terms and types appearing in - antiquotations are checked within the current theory or proof - context. - - %% FIXME less monolithic presentation, move to individual sections!? - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\isa{{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}}}[] -\rail@nont{\isa{antiquotation}}[] -\rail@term{\isa{{\isaliteral{7D}{\isacharbraceright}}}}[] -\rail@end -\rail@begin{15}{\indexdef{}{syntax}{antiquotation}\hypertarget{syntax.antiquotation}{\hyperlink{syntax.antiquotation}{\mbox{\isa{antiquotation}}}}} -\rail@bar -\rail@term{\hyperlink{antiquotation.theory}{\mbox{\isa{theory}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{antiquotation.thm}{\mbox{\isa{thm}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{antiquotation.lemma}{\mbox{\isa{lemma}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@term{\isa{\isakeyword{by}}}[] -\rail@nont{\hyperlink{syntax.method}{\mbox{\isa{method}}}}[] -\rail@bar -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.method}{\mbox{\isa{method}}}}[] -\rail@endbar -\rail@nextbar{4} -\rail@term{\hyperlink{antiquotation.prop}{\mbox{\isa{prop}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{antiquotation.term}{\mbox{\isa{term}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{6} -\rail@term{\hyperlink{antiquotation.HOL.value}{\mbox{\isa{value}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{7} -\rail@term{\hyperlink{antiquotation.term-type}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}type}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{8} -\rail@term{\hyperlink{antiquotation.typeof}{\mbox{\isa{typeof}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\isa{styles}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{9} -\rail@term{\hyperlink{antiquotation.const}{\mbox{\isa{const}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{10} -\rail@term{\hyperlink{antiquotation.abbrev}{\mbox{\isa{abbrev}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{11} -\rail@term{\hyperlink{antiquotation.typ}{\mbox{\isa{typ}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@nextbar{12} -\rail@term{\hyperlink{antiquotation.type}{\mbox{\isa{type}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{13} -\rail@term{\hyperlink{antiquotation.class}{\mbox{\isa{class}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{14} -\rail@term{\hyperlink{antiquotation.text}{\mbox{\isa{text}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\rail@begin{9}{\hyperlink{syntax.antiquotation}{\mbox{\isa{antiquotation}}}} -\rail@bar -\rail@term{\hyperlink{antiquotation.goals}{\mbox{\isa{goals}}}}[] -\rail@nont{\isa{options}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{antiquotation.subgoals}{\mbox{\isa{subgoals}}}}[] -\rail@nont{\isa{options}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{antiquotation.prf}{\mbox{\isa{prf}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{antiquotation.full-prf}{\mbox{\isa{full{\isaliteral{5F}{\isacharunderscore}}prf}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{antiquotation.ML}{\mbox{\isa{ML}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{antiquotation.ML-op}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}op}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{6} -\rail@term{\hyperlink{antiquotation.ML-type}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}type}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{7} -\rail@term{\hyperlink{antiquotation.ML-struct}{\mbox{\isa{ML{\isaliteral{5F}{\isacharunderscore}}struct}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{8} -\rail@term{\hyperlink{antiquotation.file}{\mbox{\isa{file}}}}[] -\rail@nont{\isa{options}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{options}} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\isa{option}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@endbar -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@end -\rail@begin{2}{\isa{option}} -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{styles}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\isa{style}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\isa{style}} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - Note that the syntax of antiquotations may \emph{not} include source - comments \verb|(*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*)| nor verbatim - text \verb|{|\verb|*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*|\verb|}|. - - \begin{description} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}theory\ A{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints the name \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}}, which is - guaranteed to refer to a valid ancestor theory in the current - context. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}thm\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints theorems \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. - Full fact expressions are allowed here, including attributes - (\secref{sec:syn-att}). - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}prop\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a well-typed proposition \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}lemma\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ by\ m{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} proves a well-typed proposition - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} by method \isa{m} and prints the original \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}term\ t{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a well-typed term \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}value\ t{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} evaluates a term \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} and prints - its result, see also \indexref{HOL}{command}{value}\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}term{\isaliteral{5F}{\isacharunderscore}}type\ t{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a well-typed term \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} - annotated with its type. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}typeof\ t{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints the type of a well-typed term - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}const\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a logical or syntactic constant - \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}abbrev\ c\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a constant abbreviation - \isa{{\isaliteral{22}{\isachardoublequote}}c\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ rhs{\isaliteral{22}{\isachardoublequote}}} as defined in the current context. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}typ\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a well-formed type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}type\ {\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a (logical or syntactic) type - constructor \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}class\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints a class \isa{c}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}text\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints uninterpreted source text \isa{s}. This is particularly useful to print portions of text according - to the Isabelle document style, without demanding well-formedness, - e.g.\ small pieces of terms that should not be parsed or - type-checked yet. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}goals{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints the current \emph{dynamic} goal - state. This is mainly for support of tactic-emulation scripts - within Isar. Presentation of goal states does not conform to the - idea of human-readable proof documents! - - When explaining proofs in detail it is usually better to spell out - the reasoning via proper Isar proof commands, instead of peeking at - the internal machine configuration. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}subgoals{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} is similar to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}goals{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, but - does not print the main goal. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}prf\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} prints the (compact) proof terms - corresponding to the theorems \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Note that this - requires proof terms to be switched on for the current logic - session. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}full{\isaliteral{5F}{\isacharunderscore}}prf\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} is like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}prf\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, but prints the full proof terms, i.e.\ also displays - information omitted in the compact proof term, which is denoted by - ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' placeholders there. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}ML\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}ML{\isaliteral{5F}{\isacharunderscore}}op\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}ML{\isaliteral{5F}{\isacharunderscore}}type\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}ML{\isaliteral{5F}{\isacharunderscore}}struct\ s{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} check text \isa{s} as ML value, - infix operator, type, and structure, respectively. The source is - printed verbatim. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}file\ path{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} checks that \isa{{\isaliteral{22}{\isachardoublequote}}path{\isaliteral{22}{\isachardoublequote}}} refers to a - file (or directory) and prints it verbatim. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Styled antiquotations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The antiquotations \isa{thm}, \isa{prop} and \isa{term} admit an extra \emph{style} specification to modify the - printed result. A style is specified by a name with a possibly - empty number of arguments; multiple styles can be sequenced with - commas. The following standard styles are available: - - \begin{description} - - \item \isa{lhs} extracts the first argument of any application - form with at least two arguments --- typically meta-level or - object-level equality, or any other binary relation. - - \item \isa{rhs} is like \isa{lhs}, but extracts the second - argument. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}concl{\isaliteral{22}{\isachardoublequote}}} extracts the conclusion \isa{C} from a rule - in Horn-clause normal form \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}prem{\isaliteral{22}{\isachardoublequote}}} \isa{n} extract premise number - \isa{{\isaliteral{22}{\isachardoublequote}}n{\isaliteral{22}{\isachardoublequote}}} from from a rule in Horn-clause - normal form \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{General options% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following options are available to tune the printed output - of antiquotations. Note that many of these coincide with global ML - flags of the same names. - - \begin{description} - - \item \indexdef{}{antiquotation option}{show\_types}\hypertarget{antiquotation option.show-types}{\hyperlink{antiquotation option.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} and - \indexdef{}{antiquotation option}{show\_sorts}\hypertarget{antiquotation option.show-sorts}{\hyperlink{antiquotation option.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} control - printing of explicit type and sort constraints. - - \item \indexdef{}{antiquotation option}{show\_structs}\hypertarget{antiquotation option.show-structs}{\hyperlink{antiquotation option.show-structs}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}structs}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} - controls printing of implicit structures. - - \item \indexdef{}{antiquotation option}{show\_abbrevs}\hypertarget{antiquotation option.show-abbrevs}{\hyperlink{antiquotation option.show-abbrevs}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} - controls folding of abbreviations. - - \item \indexdef{}{antiquotation option}{names\_long}\hypertarget{antiquotation option.names-long}{\hyperlink{antiquotation option.names-long}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}long}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} forces - names of types and constants etc.\ to be printed in their fully - qualified internal form. - - \item \indexdef{}{antiquotation option}{names\_short}\hypertarget{antiquotation option.names-short}{\hyperlink{antiquotation option.names-short}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}short}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} - forces names of types and constants etc.\ to be printed unqualified. - Note that internalizing the output again in the current context may - well yield a different result. - - \item \indexdef{}{antiquotation option}{names\_unique}\hypertarget{antiquotation option.names-unique}{\hyperlink{antiquotation option.names-unique}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}unique}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} - determines whether the printed version of qualified names should be - made sufficiently long to avoid overlap with names declared further - back. Set to \isa{false} for more concise output. - - \item \indexdef{}{antiquotation option}{eta\_contract}\hypertarget{antiquotation option.eta-contract}{\hyperlink{antiquotation option.eta-contract}{\mbox{\isa{eta{\isaliteral{5F}{\isacharunderscore}}contract}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} - prints terms in \isa{{\isaliteral{5C3C6574613E}{\isasymeta}}}-contracted form. - - \item \indexdef{}{antiquotation option}{display}\hypertarget{antiquotation option.display}{\hyperlink{antiquotation option.display}{\mbox{\isa{display}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} indicates - if the text is to be output as multi-line ``display material'', - rather than a small piece of text without line breaks (which is the - default). - - In this mode the embedded entities are printed in the same style as - the main theory text. - - \item \indexdef{}{antiquotation option}{break}\hypertarget{antiquotation option.break}{\hyperlink{antiquotation option.break}{\mbox{\isa{break}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} controls - line breaks in non-display material. - - \item \indexdef{}{antiquotation option}{quotes}\hypertarget{antiquotation option.quotes}{\hyperlink{antiquotation option.quotes}{\mbox{\isa{quotes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} indicates - if the output should be enclosed in double quotes. - - \item \indexdef{}{antiquotation option}{mode}\hypertarget{antiquotation option.mode}{\hyperlink{antiquotation option.mode}{\mbox{\isa{mode}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ name{\isaliteral{22}{\isachardoublequote}}} adds \isa{name} to the print mode to be used for presentation. Note that the - standard setup for {\LaTeX} output is already present by default, - including the modes \isa{latex} and \isa{xsymbols}. - - \item \indexdef{}{antiquotation option}{margin}\hypertarget{antiquotation option.margin}{\hyperlink{antiquotation option.margin}{\mbox{\isa{margin}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ nat{\isaliteral{22}{\isachardoublequote}}} and - \indexdef{}{antiquotation option}{indent}\hypertarget{antiquotation option.indent}{\hyperlink{antiquotation option.indent}{\mbox{\isa{indent}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ nat{\isaliteral{22}{\isachardoublequote}}} change the margin - or indentation for pretty printing of display material. - - \item \indexdef{}{antiquotation option}{goals\_limit}\hypertarget{antiquotation option.goals-limit}{\hyperlink{antiquotation option.goals-limit}{\mbox{\isa{goals{\isaliteral{5F}{\isacharunderscore}}limit}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ nat{\isaliteral{22}{\isachardoublequote}}} - determines the maximum number of goals to be printed (for goal-based - antiquotation). - - \item \indexdef{}{antiquotation option}{source}\hypertarget{antiquotation option.source}{\hyperlink{antiquotation option.source}{\mbox{\isa{source}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ bool{\isaliteral{22}{\isachardoublequote}}} prints the - original source text of the antiquotation arguments, rather than its - internal representation. Note that formal checking of - \hyperlink{antiquotation.thm}{\mbox{\isa{thm}}}, \hyperlink{antiquotation.term}{\mbox{\isa{term}}}, etc. is still - enabled; use the \hyperlink{antiquotation.text}{\mbox{\isa{text}}} antiquotation for unchecked - output. - - Regular \isa{{\isaliteral{22}{\isachardoublequote}}term{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}typ{\isaliteral{22}{\isachardoublequote}}} antiquotations with \isa{{\isaliteral{22}{\isachardoublequote}}source\ {\isaliteral{3D}{\isacharequal}}\ false{\isaliteral{22}{\isachardoublequote}}} involve a full round-trip from the original source - to an internalized logical entity back to a source form, according - to the syntax of the current context. Thus the printed output is - not under direct control of the author, it may even fluctuate a bit - as the underlying theory is changed later on. - - In contrast, \hyperlink{antiquotation option.source}{\mbox{\isa{source}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}\ true{\isaliteral{22}{\isachardoublequote}}} - admits direct printing of the given source text, with the desirable - well-formedness check in the background, but without modification of - the printed text. - - \end{description} - - For boolean flags, ``\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ true{\isaliteral{22}{\isachardoublequote}}}'' may be abbreviated as - ``\isa{name}''. All of the above flags are disabled by default, - unless changed from ML, say in the \verb|ROOT.ML| of the - logic session.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Markup via command tags \label{sec:tags}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Each Isabelle/Isar command may be decorated by additional - presentation tags, to indicate some modification in the way it is - printed in the document. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{tags}\hypertarget{syntax.tags}{\hyperlink{syntax.tags}{\mbox{\isa{tags}}}}} -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{tag}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{tag}} -\rail@term{\isa{{\isaliteral{25}{\isacharpercent}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.ident}{\mbox{\isa{ident}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - Some tags are pre-declared for certain classes of commands, serving - as default markup if no tags are given in the text: - - \medskip - \begin{tabular}{ll} - \isa{{\isaliteral{22}{\isachardoublequote}}theory{\isaliteral{22}{\isachardoublequote}}} & theory begin/end \\ - \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{22}{\isachardoublequote}}} & all proof commands \\ - \isa{{\isaliteral{22}{\isachardoublequote}}ML{\isaliteral{22}{\isachardoublequote}}} & all commands involving ML code \\ - \end{tabular} - - \medskip The Isabelle document preparation system - \cite{isabelle-sys} allows tagged command regions to be presented - specifically, e.g.\ to fold proof texts, or drop parts of the text - completely. - - For example ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{25}{\isacharpercent}}invisible\ auto{\isaliteral{22}{\isachardoublequote}}}'' causes - that piece of proof to be treated as \isa{invisible} instead of - \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{22}{\isachardoublequote}}} (the default), which may be shown or hidden - depending on the document setup. In contrast, ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{25}{\isacharpercent}}visible\ auto{\isaliteral{22}{\isachardoublequote}}}'' forces this text to be shown - invariably. - - Explicit tag specifications within a proof apply to all subsequent - commands of the same level of nesting. For example, ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{25}{\isacharpercent}}visible\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}'' forces the whole - sub-proof to be typeset as \isa{visible} (unless some of its parts - are tagged differently). - - \medskip Command tags merely produce certain markup environments for - type-setting. The meaning of these is determined by {\LaTeX} - macros, as defined in \verb|~~/lib/texinputs/isabelle.sty| or - by the document author. The Isabelle document preparation tools - also provide some high-level options to specify the meaning of - arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding - parts of the text. Logic sessions may also specify ``document - versions'', where given tags are interpreted in some particular way. - Again see \cite{isabelle-sys} for further details.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Railroad diagrams% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{antiquotation}{rail}\hypertarget{antiquotation.rail}{\hyperlink{antiquotation.rail}{\mbox{\isa{rail}}}} & : & \isa{antiquotation} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\isa{rail}}[] -\rail@nont{\isa{string}}[] -\rail@end -\end{railoutput} - - - The \hyperlink{antiquotation.rail}{\mbox{\isa{rail}}} antiquotation allows to include syntax - diagrams into Isabelle documents. {\LaTeX} requires the style file - \verb|~~/lib/texinputs/pdfsetup.sty|, which can be used via - \verb|\usepackage{pdfsetup}| in \verb|root.tex|, for - example. - - The rail specification language is quoted here as Isabelle \hyperlink{syntax.string}{\mbox{\isa{string}}}; it has its own grammar given below. - - \begin{railoutput} -\rail@begin{3}{} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{3B}{\isacharsemicolon}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{rule}} -\rail@bar -\rail@nextbar{1} -\rail@bar -\rail@nont{\isa{identifier}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.antiquotation}{\mbox{\isa{antiquotation}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\isa{body}}[] -\rail@end -\rail@begin{2}{\isa{body}} -\rail@plus -\rail@nont{\isa{concatenation}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{concatenation}} -\rail@plus -\rail@nont{\isa{atom}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@bar -\rail@term{\isa{{\isaliteral{2A}{\isacharasterisk}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{2} -\rail@nont{\isa{atom}}[] -\rail@endbar -\rail@endbar -\rail@end -\rail@begin{6}{\isa{atom}} -\rail@bar -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{body}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextbar{2} -\rail@nont{\isa{identifier}}[] -\rail@nextbar{3} -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{40}{\isacharat}}}}[] -\rail@endbar -\rail@bar -\rail@nont{\isa{string}}[] -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.antiquotation}{\mbox{\isa{antiquotation}}}}[] -\rail@endbar -\rail@nextbar{5} -\rail@term{\isa{{\isaliteral{5C}{\isacharbackslash}}{\isaliteral{5C}{\isacharbackslash}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - The lexical syntax of \isa{{\isaliteral{22}{\isachardoublequote}}identifier{\isaliteral{22}{\isachardoublequote}}} coincides with that of - \hyperlink{syntax.ident}{\mbox{\isa{ident}}} in regular Isabelle syntax, but \isa{string} uses - single quotes instead of double quotes of the standard \hyperlink{syntax.string}{\mbox{\isa{string}}} category, to avoid extra escapes. - - Each \isa{rule} defines a formal language (with optional name), - using a notation that is similar to EBNF or regular expressions with - recursion. The meaning and visual appearance of these rail language - elements is illustrated by the following representative examples. - - \begin{itemize} - - \item Empty \verb|()| - - \begin{railoutput} -\rail@begin{1}{} -\rail@end -\end{railoutput} - - - \item Nonterminal \verb|A| - - \begin{railoutput} -\rail@begin{1}{} -\rail@nont{\isa{A}}[] -\rail@end -\end{railoutput} - - - \item Nonterminal via Isabelle antiquotation - \verb|@{syntax method}| - - \begin{railoutput} -\rail@begin{1}{} -\rail@nont{\hyperlink{syntax.method}{\mbox{\isa{method}}}}[] -\rail@end -\end{railoutput} - - - \item Terminal \verb|'xyz'| - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\isa{xyz}}[] -\rail@end -\end{railoutput} - - - \item Terminal in keyword style \verb|@'xyz'| - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\isa{\isakeyword{xyz}}}[] -\rail@end -\end{railoutput} - - - \item Terminal via Isabelle antiquotation - \verb|@@{method rule}| - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{method.rule}{\mbox{\isa{rule}}}}[] -\rail@end -\end{railoutput} - - - \item Concatenation \verb|A B C| - - \begin{railoutput} -\rail@begin{1}{} -\rail@nont{\isa{A}}[] -\rail@nont{\isa{B}}[] -\rail@nont{\isa{C}}[] -\rail@end -\end{railoutput} - - - \item Linebreak \verb|\\| inside - concatenation\footnote{Strictly speaking, this is only a single - backslash, but the enclosing \hyperlink{syntax.string}{\mbox{\isa{string}}} syntax requires a - second one for escaping.} \verb|A B C \\ D E F| - - \begin{railoutput} -\rail@begin{3}{} -\rail@nont{\isa{A}}[] -\rail@nont{\isa{B}}[] -\rail@nont{\isa{C}}[] -\rail@cr{2} -\rail@nont{\isa{D}}[] -\rail@nont{\isa{E}}[] -\rail@nont{\isa{F}}[] -\rail@end -\end{railoutput} - - - \item Variants \verb|A |\verb,|,\verb| B |\verb,|,\verb| C| - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@nont{\isa{A}}[] -\rail@nextbar{1} -\rail@nont{\isa{B}}[] -\rail@nextbar{2} -\rail@nont{\isa{C}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \item Option \verb|A ?| - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{A}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \item Repetition \verb|A *| - - \begin{railoutput} -\rail@begin{2}{} -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{A}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \item Repetition with separator \verb|A * sep| - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\isa{A}}[] -\rail@nextplus{2} -\rail@cnont{\isa{sep}}[] -\rail@endplus -\rail@endbar -\rail@end -\end{railoutput} - - - \item Strict repetition \verb|A +| - - \begin{railoutput} -\rail@begin{2}{} -\rail@plus -\rail@nont{\isa{A}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - \item Strict repetition with separator \verb|A + sep| - - \begin{railoutput} -\rail@begin{2}{} -\rail@plus -\rail@nont{\isa{A}}[] -\rail@nextplus{1} -\rail@cnont{\isa{sep}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Draft presentation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{display\_drafts}\hypertarget{command.display-drafts}{\hyperlink{command.display-drafts}{\mbox{\isa{\isacommand{display{\isaliteral{5F}{\isacharunderscore}}drafts}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_drafts}\hypertarget{command.print-drafts}{\hyperlink{command.print-drafts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}drafts}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.display-drafts}{\mbox{\isa{\isacommand{display{\isaliteral{5F}{\isacharunderscore}}drafts}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.print-drafts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}drafts}}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.display-drafts}{\mbox{\isa{\isacommand{display{\isaliteral{5F}{\isacharunderscore}}drafts}}}}~\isa{paths} and \hyperlink{command.print-drafts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}drafts}}}}~\isa{paths} perform simple output of a given list - of raw source files. Only those symbols that do not require - additional {\LaTeX} packages are displayed properly, everything else - is left verbatim. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/First_Order_Logic.tex --- a/doc-src/IsarRef/Thy/document/First_Order_Logic.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1417 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{First{\isaliteral{5F}{\isacharunderscore}}Order{\isaliteral{5F}{\isacharunderscore}}Logic}% -% -\isamarkupheader{Example: First-Order Logic% -} -\isamarkuptrue% -% -\isadelimvisible -% -\endisadelimvisible -% -\isatagvisible -\isacommand{theory}\isamarkupfalse% -\ First{\isaliteral{5F}{\isacharunderscore}}Order{\isaliteral{5F}{\isacharunderscore}}Logic\isanewline -\isakeyword{imports}\ Base\ \ \isanewline -\isakeyword{begin}% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -% -\begin{isamarkuptext}% -\noindent In order to commence a new object-logic within - Isabelle/Pure we introduce abstract syntactic categories \isa{{\isaliteral{22}{\isachardoublequote}}i{\isaliteral{22}{\isachardoublequote}}} - for individuals and \isa{{\isaliteral{22}{\isachardoublequote}}o{\isaliteral{22}{\isachardoublequote}}} for object-propositions. The latter - is embedded into the language of Pure propositions by means of a - separate judgment.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedecl}\isamarkupfalse% -\ i\isanewline -\isacommand{typedecl}\isamarkupfalse% -\ o\isanewline -\isanewline -\isacommand{judgment}\isamarkupfalse% -\isanewline -\ \ Trueprop\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequoteclose}}\ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isaliteral{29}{\isacharparenright}}% -\begin{isamarkuptext}% -\noindent Note that the object-logic judgement is implicit in the - syntax: writing \isa{A} produces \isa{{\isaliteral{22}{\isachardoublequote}}Trueprop\ A{\isaliteral{22}{\isachardoublequote}}} internally. - From the Pure perspective this means ``\isa{A} is derivable in the - object-logic''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Equational reasoning \label{sec:framework-ex-equal}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Equality is axiomatized as a binary predicate on individuals, with - reflexivity as introduction, and substitution as elimination - principle. Note that the latter is particularly convenient in a - framework like Isabelle, because syntactic congruences are - implicitly produced by unification of \isa{{\isaliteral{22}{\isachardoublequote}}B\ x{\isaliteral{22}{\isachardoublequote}}} against - expressions containing occurrences of \isa{x}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ equal\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infix}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline -\isakeyword{where}\isanewline -\ \ refl\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ subst\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -\noindent Substitution is very powerful, but also hard to control in - full generality. We derive some common symmetry~/ transitivity - schemes of \isa{equal} as particular consequences.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ sym\ {\isaliteral{5B}{\isacharbrackleft}}sym{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ forw{\isaliteral{5F}{\isacharunderscore}}subst\ {\isaliteral{5B}{\isacharbrackleft}}trans{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ this\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ back{\isaliteral{5F}{\isacharunderscore}}subst\ {\isaliteral{5B}{\isacharbrackleft}}trans{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ trans\ {\isaliteral{5B}{\isacharbrackleft}}trans{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}y\ {\isaliteral{3D}{\isacharequal}}\ z{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ z{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Basic group theory% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As an example for equational reasoning we consider some bits of - group theory. The subsequent locale definition postulates group - operations and axioms; we also derive some consequences of this - specification.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ group\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \isakeyword{fixes}\ prod\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infix}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C636972633E}{\isasymcirc}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{7}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isakeyword{and}\ inv\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ i{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isadigit{1}}{\isadigit{0}}{\isadigit{0}}{\isadigit{0}}{\isaliteral{5D}{\isacharbrackright}}\ {\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isakeyword{and}\ unit\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ i\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isakeyword{assumes}\ assoc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ z\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isakeyword{and}\ left{\isaliteral{5F}{\isacharunderscore}}unit{\isaliteral{3A}{\isacharcolon}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isakeyword{and}\ left{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ right{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}unit\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ assoc\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}inv\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ assoc{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ assoc{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}unit{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ right{\isaliteral{5F}{\isacharunderscore}}unit{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}inv\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ assoc\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ right{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ left{\isaliteral{5F}{\isacharunderscore}}unit{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Reasoning from basic axioms is often tedious. Our proofs - work by producing various instances of the given rules (potentially - the symmetric form) using the pattern ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{eq}~\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}rule\ r{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' and composing the chain of - results via \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}/\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}. These steps may - involve any of the transitivity rules declared in - \secref{sec:framework-ex-equal}, namely \isa{trans} in combining - the first two results in \isa{right{\isaliteral{5F}{\isacharunderscore}}inv} and in the final steps of - both proofs, \isa{forw{\isaliteral{5F}{\isacharunderscore}}subst} in the first combination of \isa{right{\isaliteral{5F}{\isacharunderscore}}unit}, and \isa{back{\isaliteral{5F}{\isacharunderscore}}subst} in all other calculational steps. - - Occasional substitutions in calculations are adequate, but should - not be over-emphasized. The other extreme is to compose a chain by - plain transitivity only, with replacements occurring always in - topmost position. For example:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{unfolding}\isamarkupfalse% -\ left{\isaliteral{5F}{\isacharunderscore}}inv\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{5C3C696E76657273653E}{\isasyminverse}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{unfolding}\isamarkupfalse% -\ assoc\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{unfolding}\isamarkupfalse% -\ right{\isaliteral{5F}{\isacharunderscore}}inv\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{unfolding}\isamarkupfalse% -\ left{\isaliteral{5F}{\isacharunderscore}}unit\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Here we have re-used the built-in mechanism for unfolding - definitions in order to normalize each equational problem. A more - realistic object-logic would include proper setup for the Simplifier - (\secref{sec:simplifier}), the main automated tool for equational - reasoning in Isabelle. Then ``\hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{left{\isaliteral{5F}{\isacharunderscore}}inv}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' would become ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}simp\ only{\isaliteral{3A}{\isacharcolon}}\ left{\isaliteral{5F}{\isacharunderscore}}inv{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' etc.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Propositional logic \label{sec:framework-ex-prop}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We axiomatize basic connectives of propositional logic: implication, - disjunction, and conjunction. The associated rules are modeled - after Gentzen's system of Natural Deduction \cite{Gentzen:1935}.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ imp\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixr}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{2}}{\isadigit{5}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ impI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ impD\ {\isaliteral{5B}{\isacharbrackleft}}dest{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ disj\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixr}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{3}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ disjI\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ disjI\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ disjE\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ conj\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infixr}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{3}}{\isadigit{5}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ conjI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ conjD\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ conjD\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -\noindent The conjunctive destructions have the disadvantage that - decomposing \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequote}}} involves an immediate decision which - component should be projected. The more convenient simultaneous - elimination \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} can be derived as - follows:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\ conjE\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{obtains}\ A\ \isakeyword{and}\ B\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ conjD\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ conjD\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Here is an example of swapping conjuncts with a single - intermediate elimination step:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ B\ \isakeyword{and}\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Note that the analogous elimination rule for disjunction - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B\ {\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ A\ {\isaliteral{5C3C424241523E}{\isasymBBAR}}\ B{\isaliteral{22}{\isachardoublequote}}}'' coincides with - the original axiomatization of \isa{disjE}. - - \medskip We continue propositional logic by introducing absurdity - with its characteristic elimination. Plain truth may then be - defined as a proposition that is trivially true.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ false\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ o\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ falseE\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ true\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ o\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C746F703E}{\isasymtop}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C746F703E}{\isasymtop}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ trueI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C746F703E}{\isasymtop}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ true{\isaliteral{5F}{\isacharunderscore}}def\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\medskip\noindent Now negation represents an implication towards - absurdity:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{definition}\isamarkupfalse% -\isanewline -\ \ not\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isadigit{4}}{\isadigit{0}}{\isaliteral{5D}{\isacharbrackright}}\ {\isadigit{4}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ notI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{unfolding}\isamarkupfalse% -\ not{\isaliteral{5F}{\isacharunderscore}}def\isanewline -\isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{29}{\isacharparenright}}\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ notE\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A\isanewline -\ \ \isakeyword{shows}\ B\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{unfolding}\isamarkupfalse% -\ not{\isaliteral{5F}{\isacharunderscore}}def\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Classical logic% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Subsequently we state the principle of classical contradiction as a - local assumption. Thus we refrain from forcing the object-logic - into the classical perspective. Within that context, we may derive - well-known consequences of the classical principle.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{locale}\isamarkupfalse% -\ classical\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \isakeyword{assumes}\ classical{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{begin}\isanewline -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ double{\isaliteral{5F}{\isacharunderscore}}negation{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ C\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ classical{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{with}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{theorem}\isamarkupfalse% -\ tertium{\isaliteral{5F}{\isacharunderscore}}non{\isaliteral{5F}{\isacharunderscore}}datur{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ double{\isaliteral{5F}{\isacharunderscore}}negation{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% -\ C\ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{with}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{with}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}C\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent These examples illustrate both classical reasoning and - non-trivial propositional proofs in general. All three rules - characterize classical logic independently, but the original rule is - already the most convenient to use, because it leaves the conclusion - unchanged. Note that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ C\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} fits again into our - format for eliminations, despite the additional twist that the - context refers to the main conclusion. So we may write \isa{classical} as the Isar statement ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ thesis{\isaliteral{22}{\isachardoublequote}}}''. - This also explains nicely how classical reasoning really works: - whatever the main \isa{thesis} might be, we may always assume its - negation!% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Quantifiers \label{sec:framework-ex-quant}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Representing quantifiers is easy, thanks to the higher-order nature - of the underlying framework. According to the well-known technique - introduced by Church \cite{church40}, quantifiers are operators on - predicates, which are syntactically represented as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms - of type \isa{{\isaliteral{22}{\isachardoublequote}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequote}}}. Binder notation turns \isa{{\isaliteral{22}{\isachardoublequote}}All\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} into \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequote}}} etc.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ All\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{binder}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{1}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ allI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ allD\ {\isaliteral{5B}{\isacharbrackleft}}dest{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{axiomatization}\isamarkupfalse% -\isanewline -\ \ Ex\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{binder}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{1}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\ \isakeyword{where}\isanewline -\ \ exI\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ exE\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -\noindent The statement of \isa{exE} corresponds to ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ x\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ B\ x{\isaliteral{22}{\isachardoublequote}}}'' in Isar. In the - subsequent example we illustrate quantifier reasoning involving all - four rules:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{theorem}\isamarkupfalse% -\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ \ \ \ % -\isamarkupcmt{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}} introduction% -} -\isanewline -\ \ \isacommand{obtain}\isamarkupfalse% -\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{using}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\ \ \ \ % -\isamarkupcmt{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}} elimination% -} -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ y\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}R\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{using}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\ \ \ \ % -\isamarkupcmt{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}} destruction% -} -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ R\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\ \ \ \ % -\isamarkupcmt{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}} introduction% -} -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Canonical reasoning patterns% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The main rules of first-order predicate logic from - \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} - can now be summarized as follows, using the native Isar statement - format of \secref{sec:framework-stmt}. - - \medskip - \begin{tabular}{l} - \isa{{\isaliteral{22}{\isachardoublequote}}impI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}impD{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ A\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ B{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}disjI\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}disjI\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ B\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}disjE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B\ {\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ A\ {\isaliteral{5C3C424241523E}{\isasymBBAR}}\ B{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}conjI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ B\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}conjE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ A\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ B{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}falseE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ A{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}trueI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ {\isaliteral{5C3C746F703E}{\isasymtop}}{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}notI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}notE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ A\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ B{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}allI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}allE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ B\ a{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}exI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ B\ a\ {\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}exE{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ a\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ B\ a{\isaliteral{22}{\isachardoublequote}}} - \end{tabular} - \medskip - - \noindent This essentially provides a declarative reading of Pure - rules as Isar reasoning patterns: the rule statements tells how a - canonical proof outline shall look like. Since the above rules have - already been declared as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}} --- each according to its - particular shape --- we can immediately write Isar proof texts as - follows:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ B\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isakeyword{and}\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C746F703E}{\isasymtop}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth} -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\end{minipage} -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\bigskip\noindent Of course, these proofs are merely examples. As - sketched in \secref{sec:framework-subproof}, there is a fair amount - of flexibility in expressing Pure deductions in Isar. Here the user - is asked to express himself adequately, aiming at proof texts of - literary quality.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimvisible -% -\endisadelimvisible -% -\isatagvisible -\isacommand{end}\isamarkupfalse% -% -\endisatagvisible -{\isafoldvisible}% -% -\isadelimvisible -% -\endisadelimvisible -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Framework.tex --- a/doc-src/IsarRef/Thy/document/Framework.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1518 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Framework}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Framework\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{The Isabelle/Isar Framework \label{ch:isar-framework}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/Isar - \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift} - is intended as a generic framework for developing formal - mathematical documents with full proof checking. Definitions and - proofs are organized as theories. An assembly of theory sources may - be presented as a printed document; see also - \chref{ch:document-prep}. - - The main objective of Isar is the design of a human-readable - structured proof language, which is called the ``primary proof - format'' in Isar terminology. Such a primary proof language is - somewhere in the middle between the extremes of primitive proof - objects and actual natural language. In this respect, Isar is a bit - more formalistic than Mizar - \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar}, - using logical symbols for certain reasoning schemes where Mizar - would prefer English words; see \cite{Wenzel-Wiedijk:2002} for - further comparisons of these systems. - - So Isar challenges the traditional way of recording informal proofs - in mathematical prose, as well as the common tendency to see fully - formal proofs directly as objects of some logical calculus (e.g.\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms in a version of type theory). In fact, Isar is - better understood as an interpreter of a simple block-structured - language for describing the data flow of local facts and goals, - interspersed with occasional invocations of proof methods. - Everything is reduced to logical inferences internally, but these - steps are somewhat marginal compared to the overall bookkeeping of - the interpretation process. Thanks to careful design of the syntax - and semantics of Isar language elements, a formal record of Isar - instructions may later appear as an intelligible text to the - attentive reader. - - The Isar proof language has emerged from careful analysis of some - inherent virtues of the existing logical framework of Isabelle/Pure - \cite{paulson-found,paulson700}, notably composition of higher-order - natural deduction rules, which is a generalization of Gentzen's - original calculus \cite{Gentzen:1935}. The approach of generic - inference systems in Pure is continued by Isar towards actual proof - texts. - - Concrete applications require another intermediate layer: an - object-logic. Isabelle/HOL \cite{isa-tutorial} (simply-typed - set-theory) is being used most of the time; Isabelle/ZF - \cite{isabelle-ZF} is less extensively developed, although it would - probably fit better for classical mathematics. - - \medskip In order to illustrate natural deduction in Isar, we shall - refer to the background theory and library of Isabelle/HOL. This - includes common notions of predicate logic, naive set-theory etc.\ - using fairly standard mathematical notation. From the perspective - of generic natural deduction there is nothing special about the - logical connectives of HOL (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}}, etc.), only the resulting reasoning principles are - relevant to the user. There are similar rules available for - set-theory operators (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E7465723E}{\isasyminter}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C756E696F6E3E}{\isasymunion}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{22}{\isachardoublequote}}}, etc.), or any other theory developed in the library (lattice - theory, topology etc.). - - Subsequently we briefly review fragments of Isar proof texts - corresponding directly to such general deduction schemes. The - examples shall refer to set-theory, to minimize the danger of - understanding connectives of predicate logic as something special. - - \medskip The following deduction performs \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E7465723E}{\isasyminter}}{\isaliteral{22}{\isachardoublequote}}}-introduction, - working forwards from assumptions towards the conclusion. We give - both the Isar text, and depict the primitive rule involved, as - determined by unification of the problem against rules that are - declared in the library context.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\medskip\begin{minipage}{0.6\textwidth} -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage}\begin{minipage}{0.4\textwidth} -% -\begin{isamarkuptext}% -\infer{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequote}}}}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\end{minipage} -% -\begin{isamarkuptext}% -\medskip\noindent Note that \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} augments the proof - context, \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} indicates that the current fact shall be - used in the next step, and \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} states an intermediate - goal. The two dots ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' refer to a complete proof of - this claim, using the indicated facts and a canonical rule from the - context. We could have been more explicit here by spelling out the - final proof step via the \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}} command:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ IntI{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent The format of the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E7465723E}{\isasyminter}}{\isaliteral{22}{\isachardoublequote}}}-introduction rule represents - the most basic inference, which proceeds from given premises to a - conclusion, without any nested proof context involved. - - The next example performs backwards introduction on \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}}, - the intersection of all sets within a given set. This requires a - nested proof of set membership within a local context, where \isa{A} is an arbitrary-but-fixed member of the collection:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\medskip\begin{minipage}{0.6\textwidth} -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -% -\endisadelimnoproof -\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage}\begin{minipage}{0.4\textwidth} -% -\begin{isamarkuptext}% -\infer{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}}}{\infer*{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5B}{\isacharbrackleft}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}}}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\end{minipage} -% -\begin{isamarkuptext}% -\medskip\noindent This Isar reasoning pattern again refers to the - primitive rule depicted above. The system determines it in the - ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}'' step, which could have been spelt out more - explicitly as ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}rule\ InterI{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}''. Note - that the rule involves both a local parameter \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}} and an - assumption \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}} in the nested reasoning. This kind of - compound rule typically demands a genuine sub-proof in Isar, working - backwards rather than forwards as seen before. In the proof body we - encounter the \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} - outline of nested sub-proofs that is typical for Isar. The final - \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} is like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} followed by an additional - refinement of the enclosing claim, using the rule derived from the - proof body. - - \medskip The next example involves \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}}, which can be - characterized as the set of all \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} such that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}A{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}}. The elimination rule for \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}} does - not mention \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequote}}} at all, but admits to obtain - directly a local \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}} such that \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}} hold. This corresponds to the following Isar proof and - inference rule, respectively:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\medskip\begin{minipage}{0.6\textwidth} -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ C% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -% -\endisadelimnoproof -\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage}\begin{minipage}{0.4\textwidth} -% -\begin{isamarkuptext}% -\infer{\isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}} & \infer*{\isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}}~}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}A{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5B}{\isacharbrackleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{2C}{\isacharcomma}}\ A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}}}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\end{minipage} -% -\begin{isamarkuptext}% -\medskip\noindent Although the Isar proof follows the natural - deduction rule closely, the text reads not as natural as - anticipated. There is a double occurrence of an arbitrary - conclusion \isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}}, which represents the final result, but is - irrelevant for now. This issue arises for any elimination rule - involving local parameters. Isar provides the derived language - element \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}, which is able to perform the same - elimination proof more conveniently:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ A\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent Here we avoid to mention the final conclusion \isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}} - and return to plain forward reasoning. The rule involved in the - ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' proof is the same as before.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{The Pure framework \label{sec:framework-pure}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Pure logic \cite{paulson-found,paulson700} is an intuitionistic - fragment of higher-order logic \cite{church40}. In type-theoretic - parlance, there are three levels of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus with - corresponding arrows \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}{\isaliteral{22}{\isachardoublequote}}}/\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}}/\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}: - - \medskip - \begin{tabular}{ll} - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}} & syntactic function space (terms depending on terms) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & universal quantification (proofs depending on terms) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}} & implication (proofs depending on proofs) \\ - \end{tabular} - \medskip - - \noindent Here only the types of syntactic terms, and the - propositions of proof terms have been shown. The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-structure of proofs can be recorded as an optional feature of - the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but - the formal system can never depend on them due to \emph{proof - irrelevance}. - - On top of this most primitive layer of proofs, Pure implements a - generic calculus for nested natural deduction rules, similar to - \cite{Schroeder-Heister:1984}. Here object-logic inferences are - internalized as formulae over \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}. - Combining such rule statements may involve higher-order unification - \cite{paulson-natural}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Primitive inferences% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Term syntax provides explicit notation for abstraction \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and application \isa{{\isaliteral{22}{\isachardoublequote}}b\ a{\isaliteral{22}{\isachardoublequote}}}, while types are usually - implicit thanks to type-inference; terms of type \isa{{\isaliteral{22}{\isachardoublequote}}prop{\isaliteral{22}{\isachardoublequote}}} are - called propositions. Logical statements are composed via \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{2E}{\isachardot}}\ B{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}. Primitive reasoning operates on - judgments of the form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}, with standard introduction - and elimination rules for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} that refer to - fixed parameters \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E697375623E}{}\isactrlisub m{\isaliteral{22}{\isachardoublequote}}} and hypotheses - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} from the context \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}}; - the corresponding proof terms are left implicit. The subsequent - inference rules define \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} inductively, relative to a - collection of axioms: - - \[ - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A{\isaliteral{22}{\isachardoublequote}}}}{(\isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}} \text{~axiom})} - \qquad - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A{\isaliteral{22}{\isachardoublequote}}}}{} - \] - - \[ - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C6E6F74696E3E}{\isasymnotin}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}}} - \qquad - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} - \] - - \[ - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{2D}{\isacharminus}}\ A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}}} - \qquad - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A{\isaliteral{22}{\isachardoublequote}}}} - \] - - Furthermore, Pure provides a built-in equality \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequote}}} with axioms for reflexivity, substitution, extensionality, - and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-conversion on \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms. - - \medskip An object-logic introduces another layer on top of Pure, - e.g.\ with types \isa{{\isaliteral{22}{\isachardoublequote}}i{\isaliteral{22}{\isachardoublequote}}} for individuals and \isa{{\isaliteral{22}{\isachardoublequote}}o{\isaliteral{22}{\isachardoublequote}}} for - propositions, term constants \isa{{\isaliteral{22}{\isachardoublequote}}Trueprop\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequote}}} as - (implicit) derivability judgment and connectives like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ o{\isaliteral{22}{\isachardoublequote}}}, and axioms for object-level - rules such as \isa{{\isaliteral{22}{\isachardoublequote}}conjI{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}allI{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequote}}}. Derived object rules are represented as theorems of - Pure. After the initial object-logic setup, further axiomatizations - are usually avoided; plain definitions and derived principles are - used exclusively.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Reasoning with rules \label{sec:framework-resolution}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Primitive inferences mostly serve foundational purposes. The main - reasoning mechanisms of Pure operate on nested natural deduction - rules expressed as formulae, using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} to bind local - parameters and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} to express entailment. Multiple - parameters and premises are represented by repeating these - connectives in a right-associative manner. - - Since \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} commute thanks to the theorem - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, we may assume w.l.o.g.\ - that rule statements always observe the normal form where - quantifiers are pulled in front of implications at each level of - nesting. This means that any Pure proposition may be presented as a - \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the - form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E697375623E}{}\isactrlisub m{\isaliteral{2E}{\isachardot}}\ H\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ H\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{22}{\isachardoublequote}}} for \isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{2C}{\isacharcomma}}\ n\ {\isaliteral{5C3C67653E}{\isasymge}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}}, and \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}} atomic, and \isa{{\isaliteral{22}{\isachardoublequote}}H\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ H\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} being recursively of the same format. - Following the convention that outermost quantifiers are implicit, - Horn clauses \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{22}{\isachardoublequote}}} are a special - case of this. - - For example, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E7465723E}{\isasyminter}}{\isaliteral{22}{\isachardoublequote}}}-introduction rule encountered before is - represented as a Pure theorem as follows: - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}IntI{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequote}}} - \] - - \noindent This is a plain Horn clause, since no further nesting on - the left is involved. The general \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{22}{\isachardoublequote}}}-introduction - corresponds to a Hereditary Harrop Formula with one additional level - of nesting: - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}InterI{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}A{\isaliteral{2E}{\isachardot}}\ A\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C413E}{\isasymA}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}{\isaliteral{5C3C413E}{\isasymA}}{\isaliteral{22}{\isachardoublequote}}} - \] - - \medskip Goals are also represented as rules: \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} states that the sub-goals \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} entail the result \isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}}; for \isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} the - goal is finished. To allow \isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{22}{\isachardoublequote}}} being a rule statement - itself, we introduce the protective marker \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ prop\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequote}}}, which is defined as identity and hidden from the user. We - initialize and finish goal states as follows: - - \[ - \begin{array}{c@ {\qquad}c} - \infer[(\indexdef{}{inference}{init}\hypertarget{inference.init}{\hyperlink{inference.init}{\mbox{\isa{init}}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}C\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}C{\isaliteral{22}{\isachardoublequote}}}}{} & - \infer[(\indexdef{}{inference}{finish}\hypertarget{inference.finish}{\hyperlink{inference.finish}{\mbox{\isa{finish}}}})]{\isa{C}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}C{\isaliteral{22}{\isachardoublequote}}}} - \end{array} - \] - - \noindent Goal states are refined in intermediate proof steps until - a finished form is achieved. Here the two main reasoning principles - are \hyperlink{inference.resolution}{\mbox{\isa{resolution}}}, for back-chaining a rule against a - sub-goal (replacing it by zero or more sub-goals), and \hyperlink{inference.assumption}{\mbox{\isa{assumption}}}, for solving a sub-goal (finding a short-circuit with - local assumptions). Below \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{22}{\isachardoublequote}}} stands for \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} (\isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{5C3C67653E}{\isasymge}}\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}}). - - \[ - \infer[(\indexdef{}{inference}{resolution}\hypertarget{inference.resolution}{\hyperlink{inference.resolution}{\mbox{\isa{resolution}}}})] - {\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec H\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec a\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}}} - {\begin{tabular}{rl} - \isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ \isaliteral{5C3C5E7665633E}{}\isactrlvec a{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}goal{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec H\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{27}{\isacharprime}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}goal\ unifier{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ B\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec a\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{3D}{\isacharequal}}\ B{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular}} - \] - - \medskip - - \[ - \infer[(\indexdef{}{inference}{assumption}\hypertarget{inference.assumption}{\hyperlink{inference.assumption}{\mbox{\isa{assumption}}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}C{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}}} - {\begin{tabular}{rl} - \isa{{\isaliteral{22}{\isachardoublequote}}goal{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec H\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}assm\ unifier{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{3D}{\isacharequal}}\ H\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}}~~\text{(for some~\isa{{\isaliteral{22}{\isachardoublequote}}H\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}})} \\ - \end{tabular}} - \] - - The following trace illustrates goal-oriented reasoning in - Isabelle/Pure: - - {\footnotesize - \medskip - \begin{tabular}{r@ {\quad}l} - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}init{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}resolution\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}resolution\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}assumption{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}resolution\ A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}assumption{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ A{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}finish{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - } - - Compositions of \hyperlink{inference.assumption}{\mbox{\isa{assumption}}} after \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} occurs quite often, typically in elimination steps. - Traditional Isabelle tactics accommodate this by a combined - \indexdef{}{inference}{elim\_resolution}\hypertarget{inference.elim-resolution}{\hyperlink{inference.elim-resolution}{\mbox{\isa{elim{\isaliteral{5F}{\isacharunderscore}}resolution}}}} principle. In contrast, Isar uses - a slightly more refined combination, where the assumptions to be - closed are marked explicitly, using again the protective marker - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}{\isaliteral{22}{\isachardoublequote}}}: - - \[ - \infer[(\hyperlink{inference.refinement}{\mbox{\isa{refinement}}})] - {\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec H\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec G{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec a\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}}} - {\begin{tabular}{rl} - \isa{{\isaliteral{22}{\isachardoublequote}}sub{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}proof{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec G\ \isaliteral{5C3C5E7665633E}{}\isactrlvec a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ \isaliteral{5C3C5E7665633E}{}\isactrlvec a{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}goal{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec H\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{27}{\isacharprime}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}goal\ unifier{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ B\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec a\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{3D}{\isacharequal}}\ B{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}assm\ unifiers{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ G\isaliteral{5C3C5E7375623E}{}\isactrlsub j\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec a\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{23}{\isacharhash}}H\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5C3C76617274686574613E}{\isasymvartheta}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \quad (for each marked \isa{{\isaliteral{22}{\isachardoublequote}}G\isaliteral{5C3C5E7375623E}{}\isactrlsub j{\isaliteral{22}{\isachardoublequote}}} some \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}H\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}}) \\ - \end{tabular}} - \] - - \noindent Here the \isa{{\isaliteral{22}{\isachardoublequote}}sub{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}proof{\isaliteral{22}{\isachardoublequote}}} rule stems from the - main \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} outline of - Isar (cf.\ \secref{sec:framework-subproof}): each assumption - indicated in the text results in a marked premise \isa{{\isaliteral{22}{\isachardoublequote}}G{\isaliteral{22}{\isachardoublequote}}} above. - The marking enforces resolution against one of the sub-goal's - premises. Consequently, \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} enables to fit the result of a sub-proof quite robustly into a - pending sub-goal, while maintaining a good measure of flexibility.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{The Isar proof language \label{sec:framework-isar}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Structured proofs are presented as high-level expressions for - composing entities of Pure (propositions, facts, and goals). The - Isar proof language allows to organize reasoning within the - underlying rule calculus of Pure, but Isar is not another logical - calculus! - - Isar is an exercise in sound minimalism. Approximately half of the - language is introduced as primitive, the rest defined as derived - concepts. The following grammar describes the core language - (category \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{22}{\isachardoublequote}}}), which is embedded into theory - specification elements such as \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}; see also - \secref{sec:framework-stmt} for the separate category \isa{{\isaliteral{22}{\isachardoublequote}}statement{\isaliteral{22}{\isachardoublequote}}}. - - \medskip - \begin{tabular}{rcl} - \isa{{\isaliteral{22}{\isachardoublequote}}theory{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}stmt{\isaliteral{22}{\isachardoublequote}}} & = & \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}statement\ proof\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}}~~\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{22}{\isachardoublequote}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}prfx\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}\ stmt\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - - \isa{prfx} & = & \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}facts{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}facts{\isaliteral{22}{\isachardoublequote}}} \\ - - \isa{stmt} & = & \hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}stmt\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ facts{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}term\ {\isaliteral{3D}{\isacharequal}}\ term{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}var\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}inference{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}\ name{\isaliteral{3A}{\isacharcolon}}\ props{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}~\isa{goal} \\ - \isa{goal} & = & \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - - \medskip Simultaneous propositions or facts may be separated by the - \hyperlink{keyword.and}{\mbox{\isa{\isakeyword{and}}}} keyword. - - \medskip The syntax for terms and propositions is inherited from - Pure (and the object-logic). A \isa{{\isaliteral{22}{\isachardoublequote}}pattern{\isaliteral{22}{\isachardoublequote}}} is a \isa{{\isaliteral{22}{\isachardoublequote}}term{\isaliteral{22}{\isachardoublequote}}} with schematic variables, to be bound by higher-order - matching. - - \medskip Facts may be referenced by name or proposition. For - example, the result of ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ A\ {\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}proof{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}}'' - becomes available both as \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{22}{\isachardoublequote}}} and - \isacharbackquoteopen\isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{22}{\isachardoublequote}}}\isacharbackquoteclose. Moreover, - fact expressions may involve attributes that modify either the - theorem or the background context. For example, the expression - ``\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5B}{\isacharbrackleft}}OF\ b{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' refers to the composition of two facts - according to the \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} inference of - \secref{sec:framework-resolution}, while ``\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' - declares a fact as introduction rule in the context. - - The special fact called ``\hyperlink{fact.this}{\mbox{\isa{this}}}'' always refers to the last - result, as produced by \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, or \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}. Since \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}} occurs - frequently together with \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} we provide some - abbreviations: - - \medskip - \begin{tabular}{rcl} - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{a} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{a}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} \\ - \hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}~\isa{a} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - The \isa{{\isaliteral{22}{\isachardoublequote}}method{\isaliteral{22}{\isachardoublequote}}} category is essentially a parameter and may be - populated later. Methods use the facts indicated by \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} or \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}, and then operate on the goal state. - Some basic methods are predefined: ``\hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}}'' leaves the goal - unchanged, ``\hyperlink{method.this}{\mbox{\isa{this}}}'' applies the facts as rules to the - goal, ``\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}'' applies the facts to another rule and the - result to the goal (both ``\hyperlink{method.this}{\mbox{\isa{this}}}'' and ``\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}'' - refer to \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} of - \secref{sec:framework-resolution}). The secondary arguments to - ``\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}'' may be specified explicitly as in ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}rule\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'', or picked from the context. In the latter case, the system - first tries rules declared as \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}} or - \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}}, followed by those declared as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}. - - The default method for \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} is ``\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}'' - (arguments picked from the context), for \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} it is - ``\hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}}''. Further abbreviations for terminal proof steps - are ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ method\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}'' for - ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}'', and ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' for ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}, and ``\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}'' for ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.this}{\mbox{\isa{this}}}''. The \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}} element operates - directly on the current facts and goal by applying equalities. - - \medskip Block structure can be indicated explicitly by ``\hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}}'', although the body of a sub-proof - already involves implicit nesting. In any case, \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} - jumps into the next section of a block, i.e.\ it acts like closing - an implicit block scope and opening another one; there is no direct - correspondence to subgoals here. - - The remaining elements \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}} and \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} build up - a local context (see \secref{sec:framework-context}), while - \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} refines a pending sub-goal by the rule resulting - from a nested sub-proof (see \secref{sec:framework-subproof}). - Further derived concepts will support calculational reasoning (see - \secref{sec:framework-calc}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Context elements \label{sec:framework-context}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In judgments \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} of the primitive framework, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}} - essentially acts like a proof context. Isar elaborates this idea - towards a higher-level notion, with additional information for - type-inference, term abbreviations, local facts, hypotheses etc. - - The element \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}} declares a local - parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in - results exported from the context, \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} may become anything. - The \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}inference{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}{\isaliteral{22}{\isachardoublequote}}} element provides a - general interface to hypotheses: ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}inference{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}\ A{\isaliteral{22}{\isachardoublequote}}}'' produces \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A{\isaliteral{22}{\isachardoublequote}}} locally, while the - included inference tells how to discharge \isa{A} from results - \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}} later on. There is no user-syntax for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}inference{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}{\isaliteral{22}{\isachardoublequote}}}, i.e.\ it may only occur internally when derived - commands are defined in ML. - - At the user-level, the default inference for \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} is - \hyperlink{inference.discharge}{\mbox{\isa{discharge}}} as given below. The additional variants - \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}} and \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} are defined as follows: - - \medskip - \begin{tabular}{rcl} - \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}~\isa{A} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}weak{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}discharge{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}\ A{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ a{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}expansion{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ a{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - \[ - \infer[(\indexdef{}{inference}{discharge}\hypertarget{inference.discharge}{\hyperlink{inference.discharge}{\mbox{\isa{discharge}}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{2D}{\isacharminus}}\ A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{23}{\isacharhash}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}}} - \] - \[ - \infer[(\indexdef{}{inference}{weak-discharge}\hypertarget{inference.weak-discharge}{\hyperlink{inference.weak-discharge}{\mbox{\isa{weak{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}discharge}}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{2D}{\isacharminus}}\ A\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}}} - \] - \[ - \infer[(\indexdef{}{inference}{expansion}\hypertarget{inference.expansion}{\hyperlink{inference.expansion}{\mbox{\isa{expansion}}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{2D}{\isacharminus}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B\ a{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73747275743E}{\isasymstrut}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B\ x{\isaliteral{22}{\isachardoublequote}}}} - \] - - \medskip Note that \hyperlink{inference.discharge}{\mbox{\isa{discharge}}} and \hyperlink{inference.weak-discharge}{\mbox{\isa{weak{\isaliteral{5C3C68797068656E3E}{\isasymhyphen}}discharge}}} differ in the marker for \isa{A}, which is - relevant when the result of a \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}-\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}-\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} outline is composed with a pending goal, - cf.\ \secref{sec:framework-subproof}. - - The most interesting derived context element in Isar is \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} \cite[\S5.3]{Wenzel-PhD}, which supports generalized - elimination steps in a purely forward manner. The \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} - command takes a specification of parameters \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{22}{\isachardoublequote}}} and - assumptions \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec A{\isaliteral{22}{\isachardoublequote}}} to be added to the context, together - with a proof of a case rule stating that this extension is - conservative (i.e.\ may be removed from closed results later on): - - \medskip - \begin{tabular}{l} - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}facts{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}}~~\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ \ {\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}proof{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \quad \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}case{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}thesis{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}} \\ - \qquad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{thesis} \\ - \qquad \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\ - \qquad \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{thesis}~\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}facts{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}\ {\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}proof{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} \\ - \quad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6775696C6C656D6F746C6566743E}{\isasymguillemotleft}}elimination\ case{\isaliteral{5C3C6775696C6C656D6F7472696768743E}{\isasymguillemotright}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - \[ - \infer[(\hyperlink{inference.elimination}{\mbox{\isa{elimination}}})]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}}}{ - \begin{tabular}{rl} - \isa{{\isaliteral{22}{\isachardoublequote}}case{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}thesis{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\[0.2ex] - \isa{{\isaliteral{22}{\isachardoublequote}}result{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec y\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ B{\isaliteral{22}{\isachardoublequote}}} \\[0.2ex] - \end{tabular}} - \] - - \noindent Here the name ``\isa{thesis}'' is a specific convention - for an arbitrary-but-fixed proposition; in the primitive natural - deduction rules shown before we have occasionally used \isa{C}. - The whole statement of ``\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{x}~\hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}}'' may be read as a claim that \isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}} - may be assumed for some arbitrary-but-fixed \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}}. Also note - that ``\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ B{\isaliteral{22}{\isachardoublequote}}}'' without parameters - is similar to ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ B{\isaliteral{22}{\isachardoublequote}}}'', but the - latter involves multiple sub-goals. - - \medskip The subsequent Isar proof texts explain all context - elements introduced above using the formal proof language itself. - After finishing a local proof within a block, we indicate the - exported result via \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{minipage}[t]{0.45\textwidth} -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{60}{\isacharbackquoteclose}}% -\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{60}{\isacharbackquoteclose}}% -\end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth} -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{def}\isamarkupfalse% -\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ a\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B\ a{\isaliteral{60}{\isacharbackquoteclose}}% -\end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{obtain}\isamarkupfalse% -\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B{\isaliteral{60}{\isacharbackquoteclose}}% -\end{minipage} -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\bigskip\noindent This illustrates the meaning of Isar context - elements without goals getting in between.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Structured statements \label{sec:framework-stmt}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The category \isa{{\isaliteral{22}{\isachardoublequote}}statement{\isaliteral{22}{\isachardoublequote}}} of top-level theorem specifications - is defined as follows: - - \medskip - \begin{tabular}{rcl} - \isa{{\isaliteral{22}{\isachardoublequote}}statement{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}context\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}\ conclusion{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - - \isa{{\isaliteral{22}{\isachardoublequote}}context{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C46495845533E}{\isasymFIXES}}\ vars\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ name{\isaliteral{3A}{\isacharcolon}}\ props\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - - \isa{{\isaliteral{22}{\isachardoublequote}}conclusion{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ name{\isaliteral{3A}{\isacharcolon}}\ props\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ vars\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ name{\isaliteral{3A}{\isacharcolon}}\ props\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \quad \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C424241523E}{\isasymBBAR}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - - \medskip\noindent A simple \isa{{\isaliteral{22}{\isachardoublequote}}statement{\isaliteral{22}{\isachardoublequote}}} consists of named - propositions. The full form admits local context elements followed - by the actual conclusions, such as ``\hyperlink{keyword.fixes}{\mbox{\isa{\isakeyword{fixes}}}}~\isa{x}~\hyperlink{keyword.assumes}{\mbox{\isa{\isakeyword{assumes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{keyword.shows}{\mbox{\isa{\isakeyword{shows}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}B\ x{\isaliteral{22}{\isachardoublequote}}}''. The final result emerges as a Pure rule after discharging - the context: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{22}{\isachardoublequote}}}. - - The \hyperlink{keyword.obtains}{\mbox{\isa{\isakeyword{obtains}}}} variant is another abbreviation defined - below; unlike \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} (cf.\ - \secref{sec:framework-context}) there may be several ``cases'' - separated by ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C424241523E}{\isasymBBAR}}{\isaliteral{22}{\isachardoublequote}}}'', each consisting of several - parameters (\isa{{\isaliteral{22}{\isachardoublequote}}vars{\isaliteral{22}{\isachardoublequote}}}) and several premises (\isa{{\isaliteral{22}{\isachardoublequote}}props{\isaliteral{22}{\isachardoublequote}}}). - This specifies multi-branch elimination rules. - - \medskip - \begin{tabular}{l} - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ \ \ {\isaliteral{5C3C424241523E}{\isasymBBAR}}\ \ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ \ \ {\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \quad \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C46495845533E}{\isasymFIXES}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}\isaliteral{5C3C5E7665633E}{}\isactrlvec x{\isaliteral{2E}{\isachardot}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec A\ \isaliteral{5C3C5E7665633E}{}\isactrlvec x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis\ \ {\isaliteral{5C3C414E443E}{\isasymAND}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - Presenting structured statements in such an ``open'' format usually - simplifies the subsequent proof, because the outer structure of the - problem is already laid out directly. E.g.\ consider the following - canonical patterns for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C53484F57533E}{\isasymSHOWS}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4F425441494E533E}{\isasymOBTAINS}}{\isaliteral{22}{\isachardoublequote}}}, - respectively:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{minipage}{0.5\textwidth} -\isacommand{theorem}\isamarkupfalse% -\isanewline -\ \ \isakeyword{fixes}\ x\ \isakeyword{and}\ y\isanewline -\ \ \isakeyword{assumes}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ x{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}B\ y{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage}\begin{minipage}{0.5\textwidth} -\isacommand{theorem}\isamarkupfalse% -\isanewline -\ \ \isakeyword{obtains}\ x\ \isakeyword{and}\ y\isanewline -\ \ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ b{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -\isanewline -% -\endisadelimnoproof -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage} -% -\begin{isamarkuptext}% -\medskip\noindent Here local facts \isacharbackquoteopen\isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}}\isacharbackquoteclose\ and \isacharbackquoteopen\isa{{\isaliteral{22}{\isachardoublequote}}B\ y{\isaliteral{22}{\isachardoublequote}}}\isacharbackquoteclose\ are referenced immediately; there is no - need to decompose the logical rule structure again. In the second - proof the final ``\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{thesis}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' involves the local rule case \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} for the particular instance of terms \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}b{\isaliteral{22}{\isachardoublequote}}} produced in the body.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Structured proof refinement \label{sec:framework-subproof}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -By breaking up the grammar for the Isar proof language, we may - understand a proof text as a linear sequence of individual proof - commands. These are interpreted as transitions of the Isar virtual - machine (Isar/VM), which operates on a block-structured - configuration in single steps. This allows users to write proof - texts in an incremental manner, and inspect intermediate - configurations for debugging. - - The basic idea is analogous to evaluating algebraic expressions on a - stack machine: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C63646F743E}{\isasymcdot}}\ c{\isaliteral{22}{\isachardoublequote}}} then corresponds to a sequence - of single transitions for each symbol \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{2C}{\isacharcomma}}\ a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{2B}{\isacharplus}}{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C63646F743E}{\isasymcdot}}{\isaliteral{2C}{\isacharcomma}}\ c{\isaliteral{22}{\isachardoublequote}}}. - In Isar the algebraic values are facts or goals, and the operations - are inferences. - - \medskip The Isar/VM state maintains a stack of nodes, each node - contains the local proof context, the linguistic mode, and a pending - goal (optional). The mode determines the type of transition that - may be performed next, it essentially alternates between forward and - backward reasoning, with an intermediate stage for chained facts - (see \figref{fig:isar-vm}). - - \begin{figure}[htb] - \begin{center} - \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm} - \end{center} - \caption{Isar/VM modes}\label{fig:isar-vm} - \end{figure} - - For example, in \isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} mode Isar acts like a mathematical - scratch-pad, accepting declarations like \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, and claims like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}. A goal - statement changes the mode to \isa{{\isaliteral{22}{\isachardoublequote}}prove{\isaliteral{22}{\isachardoublequote}}}, which means that we - may now refine the problem via \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}} or \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}. Then we are again in \isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} mode of a proof body, - which may issue \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} statements to solve pending - sub-goals. A concluding \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} will return to the original - \isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} mode one level upwards. The subsequent Isar/VM - trace indicates block structure, linguistic mode, goal state, and - inferences:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begingroup\footnotesize -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{minipage}[t]{0.18\textwidth} -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -% -\isadelimnoproof -\ \ \ \ \ \ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -% -\endisadelimnoproof -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\quad -\begin{minipage}[t]{0.06\textwidth} -\isa{{\isaliteral{22}{\isachardoublequote}}begin{\isaliteral{22}{\isachardoublequote}}} \\ -\\ -\\ -\isa{{\isaliteral{22}{\isachardoublequote}}begin{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}end{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}end{\isaliteral{22}{\isachardoublequote}}} \\ -\end{minipage} -\begin{minipage}[t]{0.08\textwidth} -\isa{{\isaliteral{22}{\isachardoublequote}}prove{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}prove{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}state{\isaliteral{22}{\isachardoublequote}}} \\ -\end{minipage}\begin{minipage}[t]{0.35\textwidth} -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{23}{\isacharhash}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\\ -\\ -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{23}{\isacharhash}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}} \\ -\end{minipage}\begin{minipage}[t]{0.4\textwidth} -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}init{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}resolution\ impI{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\\ -\\ -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}refinement\ {\isaliteral{23}{\isacharhash}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}finish{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ -\end{minipage} -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\endgroup -% -\begin{isamarkuptext}% -\noindent Here the \hyperlink{inference.refinement}{\mbox{\isa{refinement}}} inference from - \secref{sec:framework-resolution} mediates composition of Isar - sub-proofs nicely. Observe that this principle incorporates some - degree of freedom in proof composition. In particular, the proof - body allows parameters and assumptions to be re-ordered, or commuted - according to Hereditary Harrop Form. Moreover, context elements - that are not used in a sub-proof may be omitted altogether. For - example:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{minipage}{0.5\textwidth} -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ \isakeyword{and}\ y\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -% -\endisadelimnoproof -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\begin{minipage}{0.5\textwidth} -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ y\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isadelimnoproof -\ % -\endisadelimnoproof -% -\isatagnoproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagnoproof -{\isafoldnoproof}% -% -\isadelimnoproof -% -\endisadelimnoproof -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth} -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ y\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\end{minipage}\begin{minipage}{0.5\textwidth} -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ y\ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}C\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\end{minipage} -% -\begin{isamarkuptext}% -\medskip\noindent Such ``peephole optimizations'' of Isar texts are - practically important to improve readability, by rearranging - contexts elements according to the natural flow of reasoning in the - body, while still observing the overall scoping rules. - - \medskip This illustrates the basic idea of structured proof - processing in Isar. The main mechanisms are based on natural - deduction rule composition within the Pure framework. In - particular, there are no direct operations on goal states within the - proof body. Moreover, there is no hidden automated reasoning - involved, just plain unification.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Calculational reasoning \label{sec:framework-calc}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The existing Isar infrastructure is sufficiently flexible to support - calculational reasoning (chains of transitivity steps) as derived - concept. The generic proof elements introduced below depend on - rules declared as \hyperlink{attribute.trans}{\mbox{\isa{trans}}} in the context. It is left to - the object-logic to provide a suitable rule collection for mixed - relations of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3C}{\isacharless}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7375627365743E}{\isasymsubset}}{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}{\isaliteral{22}{\isachardoublequote}}} etc. Due to the flexibility of rule composition - (\secref{sec:framework-resolution}), substitution of equals by - equals is covered as well, even substitution of inequalities - involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD} - and \cite{Bauer-Wenzel:2001}. - - The generic calculational mechanism is based on the observation that - rules such as \isa{{\isaliteral{22}{\isachardoublequote}}trans{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{3D}{\isacharequal}}\ z\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{3D}{\isacharequal}}\ z{\isaliteral{22}{\isachardoublequote}}} - proceed from the premises towards the conclusion in a deterministic - fashion. Thus we may reason in forward mode, feeding intermediate - results into rules selected from the context. The course of - reasoning is organized by maintaining a secondary fact called - ``\hyperlink{fact.calculation}{\mbox{\isa{calculation}}}'', apart from the primary ``\hyperlink{fact.this}{\mbox{\isa{this}}}'' - already provided by the Isar primitives. In the definitions below, - \hyperlink{attribute.OF}{\mbox{\isa{OF}}} refers to \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} - (\secref{sec:framework-resolution}) with multiple rule arguments, - and \isa{{\isaliteral{22}{\isachardoublequote}}trans{\isaliteral{22}{\isachardoublequote}}} represents to a suitable rule from the context: - - \begin{matharray}{rcl} - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isaliteral{2B}{\isacharplus}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ trans\ {\isaliteral{5B}{\isacharbrackleft}}OF\ calculation\ this{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} & \equiv & \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\ - \end{matharray} - - \noindent The start of a calculation is determined implicitly in the - text: here \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} sets \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} to the current - result; any subsequent occurrence will update \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} by - combination with the next result and a transitivity rule. The - calculational sequence is concluded via \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}, where - the final result is exposed for use in a concluding claim. - - Here is a canonical proof pattern, using \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} to - establish the intermediate results:% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -\noindent The term ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' above is a special abbreviation - provided by the Isabelle/Isar syntax layer: it statically refers to - the right-hand side argument of the previous statement given in the - text. Thus it happens to coincide with relevant sub-expressions in - the calculational chain, but the exact correspondence is dependent - on the transitivity rules being involved. - - \medskip Symmetry rules such as \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequote}}} are like - transitivities with only one premise. Isar maintains a separate - rule collection declared via the \hyperlink{attribute.sym}{\mbox{\isa{sym}}} attribute, to be - used in fact expressions ``\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'', or single-step - proofs ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Generic.tex --- a/doc-src/IsarRef/Thy/document/Generic.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1926 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Generic}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Generic\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Generic tools and packages \label{ch:gen-tools}% -} -\isamarkuptrue% -% -\isamarkupsection{Configuration options \label{sec:config}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/Pure maintains a record of named configuration - options within the theory or proof context, with values of type - \verb|bool|, \verb|int|, \verb|real|, or \verb|string|. Tools may declare options in ML, and then refer to these - values (relative to the context). Thus global reference variables - are easily avoided. The user may change the value of a - configuration option by means of an associated attribute of the same - name. This form of context declaration works particularly well with - commands such as \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} or \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}} like - this:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{declare}\isamarkupfalse% -\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}show{\isaliteral{5F}{\isacharunderscore}}main{\isaliteral{5F}{\isacharunderscore}}goal\ {\isaliteral{3D}{\isacharequal}}\ false{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\isanewline -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{note}\isamarkupfalse% -\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}show{\isaliteral{5F}{\isacharunderscore}}main{\isaliteral{5F}{\isacharunderscore}}goal\ {\isaliteral{3D}{\isacharequal}}\ true{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -For historical reasons, some tools cannot take the full proof - context into account and merely refer to the background theory. - This is accommodated by configuration options being declared as - ``global'', which may not be changed within a local context. - - \begin{matharray}{rcll} - \indexdef{}{command}{print\_configs}\hypertarget{command.print-configs}{\hyperlink{command.print-configs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}configs}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{6}{} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@bar -\rail@term{\isa{true}}[] -\rail@nextbar{2} -\rail@term{\isa{false}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.int}{\mbox{\isa{int}}}}[] -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.float}{\mbox{\isa{float}}}}[] -\rail@nextbar{5} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.print-configs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}configs}}}} prints the available configuration - options, with names, types, and current values. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ value{\isaliteral{22}{\isachardoublequote}}} as an attribute expression modifies the - named option, with the syntax of the value depending on the option's - type. For \verb|bool| the default value is \isa{true}. Any - attempt to change a global option in a local context is ignored. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Basic proof tools% -} -\isamarkuptrue% -% -\isamarkupsubsection{Miscellaneous methods and attributes \label{sec:misc-meth-att}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{unfold}\hypertarget{method.unfold}{\hyperlink{method.unfold}{\mbox{\isa{unfold}}}} & : & \isa{method} \\ - \indexdef{}{method}{fold}\hypertarget{method.fold}{\hyperlink{method.fold}{\mbox{\isa{fold}}}} & : & \isa{method} \\ - \indexdef{}{method}{insert}\hypertarget{method.insert}{\hyperlink{method.insert}{\mbox{\isa{insert}}}} & : & \isa{method} \\[0.5ex] - \indexdef{}{method}{erule}\hypertarget{method.erule}{\hyperlink{method.erule}{\mbox{\isa{erule}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{drule}\hypertarget{method.drule}{\hyperlink{method.drule}{\mbox{\isa{drule}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{frule}\hypertarget{method.frule}{\hyperlink{method.frule}{\mbox{\isa{frule}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{intro}\hypertarget{method.intro}{\hyperlink{method.intro}{\mbox{\isa{intro}}}} & : & \isa{method} \\ - \indexdef{}{method}{elim}\hypertarget{method.elim}{\hyperlink{method.elim}{\mbox{\isa{elim}}}} & : & \isa{method} \\ - \indexdef{}{method}{succeed}\hypertarget{method.succeed}{\hyperlink{method.succeed}{\mbox{\isa{succeed}}}} & : & \isa{method} \\ - \indexdef{}{method}{fail}\hypertarget{method.fail}{\hyperlink{method.fail}{\mbox{\isa{fail}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{method.fold}{\mbox{\isa{fold}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.unfold}{\mbox{\isa{unfold}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.insert}{\mbox{\isa{insert}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{method.erule}{\mbox{\isa{erule}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.drule}{\mbox{\isa{drule}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.frule}{\mbox{\isa{frule}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{method.intro}{\mbox{\isa{intro}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.elim}{\mbox{\isa{elim}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.unfold}{\mbox{\isa{unfold}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} and \hyperlink{method.fold}{\mbox{\isa{fold}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} expand (or fold back) the given definitions throughout - all goals; any chained facts provided are inserted into the goal and - subject to rewriting as well. - - \item \hyperlink{method.insert}{\mbox{\isa{insert}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} inserts theorems as facts - into all goals of the proof state. Note that current facts - indicated for forward chaining are ignored. - - \item \hyperlink{method.erule}{\mbox{\isa{erule}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, \hyperlink{method.drule}{\mbox{\isa{drule}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, and \hyperlink{method.frule}{\mbox{\isa{frule}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are similar to the basic \hyperlink{method.rule}{\mbox{\isa{rule}}} - method (see \secref{sec:pure-meth-att}), but apply rules by - elim-resolution, destruct-resolution, and forward-resolution, - respectively \cite{isabelle-implementation}. The optional natural - number argument (default 0) specifies additional assumption steps to - be performed here. - - Note that these methods are improper ones, mainly serving for - experimentation and tactic script emulation. Different modes of - basic rule application are usually expressed in Isar at the proof - language level, rather than via implicit proof state manipulations. - For example, a proper single-step elimination would be done using - the plain \hyperlink{method.rule}{\mbox{\isa{rule}}} method, with forward chaining of current - facts. - - \item \hyperlink{method.intro}{\mbox{\isa{intro}}} and \hyperlink{method.elim}{\mbox{\isa{elim}}} repeatedly refine some goal - by intro- or elim-resolution, after having inserted any chained - facts. Exactly the rules given as arguments are taken into account; - this allows fine-tuned decomposition of a proof problem, in contrast - to common automated tools. - - \item \hyperlink{method.succeed}{\mbox{\isa{succeed}}} yields a single (unchanged) result; it is - the identity of the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}{\isaliteral{22}{\isachardoublequote}}}'' method combinator (cf.\ - \secref{sec:proof-meth}). - - \item \hyperlink{method.fail}{\mbox{\isa{fail}}} yields an empty result sequence; it is the - identity of the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}}'' method combinator (cf.\ - \secref{sec:proof-meth}). - - \end{description} - - \begin{matharray}{rcl} - \indexdef{}{attribute}{tagged}\hypertarget{attribute.tagged}{\hyperlink{attribute.tagged}{\mbox{\isa{tagged}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{untagged}\hypertarget{attribute.untagged}{\hyperlink{attribute.untagged}{\mbox{\isa{untagged}}}} & : & \isa{attribute} \\[0.5ex] - \indexdef{}{attribute}{THEN}\hypertarget{attribute.THEN}{\hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{unfolded}\hypertarget{attribute.unfolded}{\hyperlink{attribute.unfolded}{\mbox{\isa{unfolded}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{folded}\hypertarget{attribute.folded}{\hyperlink{attribute.folded}{\mbox{\isa{folded}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{abs\_def}\hypertarget{attribute.abs-def}{\hyperlink{attribute.abs-def}{\mbox{\isa{abs{\isaliteral{5F}{\isacharunderscore}}def}}}} & : & \isa{attribute} \\[0.5ex] - \indexdef{}{attribute}{rotated}\hypertarget{attribute.rotated}{\hyperlink{attribute.rotated}{\mbox{\isa{rotated}}}} & : & \isa{attribute} \\ - \indexdef{Pure}{attribute}{elim\_format}\hypertarget{attribute.Pure.elim-format}{\hyperlink{attribute.Pure.elim-format}{\mbox{\isa{elim{\isaliteral{5F}{\isacharunderscore}}format}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{standard}\hypertarget{attribute.standard}{\hyperlink{attribute.standard}{\mbox{\isa{standard}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{no\_vars}\hypertarget{attribute.no-vars}{\hyperlink{attribute.no-vars}{\mbox{\isa{no{\isaliteral{5F}{\isacharunderscore}}vars}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.tagged}{\mbox{\isa{tagged}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.untagged}{\mbox{\isa{untagged}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{attribute.unfolded}{\mbox{\isa{unfolded}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.folded}{\mbox{\isa{folded}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.rotated}{\mbox{\isa{rotated}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.int}{\mbox{\isa{int}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{attribute.tagged}{\mbox{\isa{tagged}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ value{\isaliteral{22}{\isachardoublequote}}} and \hyperlink{attribute.untagged}{\mbox{\isa{untagged}}}~\isa{name} add and remove \emph{tags} of some theorem. - Tags may be any list of string pairs that serve as formal comment. - The first string is considered the tag name, the second its value. - Note that \hyperlink{attribute.untagged}{\mbox{\isa{untagged}}} removes any tags of the same name. - - \item \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}~\isa{a} composes rules by resolution; it - resolves with the first premise of \isa{a} (an alternative - position may be also specified). See also \verb|RS| in - \cite{isabelle-implementation}. - - \item \hyperlink{attribute.unfolded}{\mbox{\isa{unfolded}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} and \hyperlink{attribute.folded}{\mbox{\isa{folded}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} expand and fold back again the given - definitions throughout a rule. - - \item \hyperlink{attribute.abs-def}{\mbox{\isa{abs{\isaliteral{5F}{\isacharunderscore}}def}}} turns an equation of the form \isa{{\isaliteral{22}{\isachardoublequote}}f\ x\ y\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} into \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ t{\isaliteral{22}{\isachardoublequote}}}, which ensures that \hyperlink{method.simp}{\mbox{\isa{simp}}} or \hyperlink{method.unfold}{\mbox{\isa{unfold}}} steps always expand it. This also works - for object-logic equality. - - \item \hyperlink{attribute.rotated}{\mbox{\isa{rotated}}}~\isa{n} rotate the premises of a - theorem by \isa{n} (default 1). - - \item \hyperlink{attribute.Pure.elim-format}{\mbox{\isa{elim{\isaliteral{5F}{\isacharunderscore}}format}}} turns a destruction rule into - elimination rule format, by resolving with the rule \isa{{\isaliteral{22}{\isachardoublequote}}PROP\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}PROP\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ PROP\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ PROP\ B{\isaliteral{22}{\isachardoublequote}}}. - - Note that the Classical Reasoner (\secref{sec:classical}) provides - its own version of this operation. - - \item \hyperlink{attribute.standard}{\mbox{\isa{standard}}} puts a theorem into the standard form of - object-rules at the outermost theory level. Note that this - operation violates the local proof context (including active - locales). - - \item \hyperlink{attribute.no-vars}{\mbox{\isa{no{\isaliteral{5F}{\isacharunderscore}}vars}}} replaces schematic variables by free - ones; this is mainly for tuning output of pretty printed theorems. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Low-level equational reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{subst}\hypertarget{method.subst}{\hyperlink{method.subst}{\mbox{\isa{subst}}}} & : & \isa{method} \\ - \indexdef{}{method}{hypsubst}\hypertarget{method.hypsubst}{\hyperlink{method.hypsubst}{\mbox{\isa{hypsubst}}}} & : & \isa{method} \\ - \indexdef{}{method}{split}\hypertarget{method.split}{\hyperlink{method.split}{\mbox{\isa{split}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{6}{} -\rail@term{\hyperlink{method.subst}{\mbox{\isa{subst}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{asm}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextplus{5} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{method.split}{\mbox{\isa{split}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\end{railoutput} - - - These methods provide low-level facilities for equational reasoning - that are intended for specialized applications only. Normally, - single step calculations would be performed in a structured text - (see also \secref{sec:calculation}), while the Simplifier methods - provide the canonical way for automated normalization (see - \secref{sec:simplifier}). - - \begin{description} - - \item \hyperlink{method.subst}{\mbox{\isa{subst}}}~\isa{eq} performs a single substitution step - using rule \isa{eq}, which may be either a meta or object - equality. - - \item \hyperlink{method.subst}{\mbox{\isa{subst}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}asm{\isaliteral{29}{\isacharparenright}}\ eq{\isaliteral{22}{\isachardoublequote}}} substitutes in an - assumption. - - \item \hyperlink{method.subst}{\mbox{\isa{subst}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ j{\isaliteral{29}{\isacharparenright}}\ eq{\isaliteral{22}{\isachardoublequote}}} performs several - substitutions in the conclusion. The numbers \isa{i} to \isa{j} - indicate the positions to substitute at. Positions are ordered from - the top of the term tree moving down from left to right. For - example, in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}c\ {\isaliteral{2B}{\isacharplus}}\ d{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} there are three positions - where commutativity of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} is applicable: 1 refers to \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{22}{\isachardoublequote}}}, 2 to the whole term, and 3 to \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{2B}{\isacharplus}}\ d{\isaliteral{22}{\isachardoublequote}}}. - - If the positions in the list \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ j{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} are non-overlapping - (e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}\ {\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}c\ {\isaliteral{2B}{\isacharplus}}\ d{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}) you may - assume all substitutions are performed simultaneously. Otherwise - the behaviour of \isa{subst} is not specified. - - \item \hyperlink{method.subst}{\mbox{\isa{subst}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}asm{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ j{\isaliteral{29}{\isacharparenright}}\ eq{\isaliteral{22}{\isachardoublequote}}} performs the - substitutions in the assumptions. The positions refer to the - assumptions in order from left to right. For example, given in a - goal of the form \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{28}{\isacharparenleft}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}c\ {\isaliteral{2B}{\isacharplus}}\ d{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}, position 1 of - commutativity of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} is the subterm \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{2B}{\isacharplus}}\ b{\isaliteral{22}{\isachardoublequote}}} and - position 2 is the subterm \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{2B}{\isacharplus}}\ d{\isaliteral{22}{\isachardoublequote}}}. - - \item \hyperlink{method.hypsubst}{\mbox{\isa{hypsubst}}} performs substitution using some - assumption; this only works for equations of the form \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} where \isa{x} is a free or bound variable. - - \item \hyperlink{method.split}{\mbox{\isa{split}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} performs single-step case - splitting using the given rules. Splitting is performed in the - conclusion or some assumption of the subgoal, depending of the - structure of the rule. - - Note that the \hyperlink{method.simp}{\mbox{\isa{simp}}} method already involves repeated - application of split rules as declared in the current context, using - \hyperlink{attribute.split}{\mbox{\isa{split}}}, for example. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Further tactic emulations \label{sec:tactics}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following improper proof methods emulate traditional tactics. - These admit direct access to the goal state, which is normally - considered harmful! In particular, this may involve both numbered - goal addressing (default 1), and dynamic instantiation within the - scope of some subgoal. - - \begin{warn} - Dynamic instantiations refer to universally quantified parameters - of a subgoal (the dynamic context) rather than fixed variables and - term abbreviations of a (static) Isar context. - \end{warn} - - Tactic emulation methods, unlike their ML counterparts, admit - simultaneous instantiation from both dynamic and static contexts. - If names occur in both contexts goal parameters hide locally fixed - variables. Likewise, schematic variables refer to term - abbreviations, if present in the static context. Otherwise the - schematic variable is interpreted as a schematic variable and left - to be solved by unification with certain parts of the subgoal. - - Note that the tactic emulation proof methods in Isabelle/Isar are - consistently named \isa{foo{\isaliteral{5F}{\isacharunderscore}}tac}. Note also that variable names - occurring on left hand sides of instantiations must be preceded by a - question mark if they coincide with a keyword or contain dots. This - is consistent with the attribute \hyperlink{attribute.where}{\mbox{\isa{where}}} (see - \secref{sec:pure-meth-att}). - - \begin{matharray}{rcl} - \indexdef{}{method}{rule\_tac}\hypertarget{method.rule-tac}{\hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{erule\_tac}\hypertarget{method.erule-tac}{\hyperlink{method.erule-tac}{\mbox{\isa{erule{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{drule\_tac}\hypertarget{method.drule-tac}{\hyperlink{method.drule-tac}{\mbox{\isa{drule{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{frule\_tac}\hypertarget{method.frule-tac}{\hyperlink{method.frule-tac}{\mbox{\isa{frule{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{cut\_tac}\hypertarget{method.cut-tac}{\hyperlink{method.cut-tac}{\mbox{\isa{cut{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{thin\_tac}\hypertarget{method.thin-tac}{\hyperlink{method.thin-tac}{\mbox{\isa{thin{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{subgoal\_tac}\hypertarget{method.subgoal-tac}{\hyperlink{method.subgoal-tac}{\mbox{\isa{subgoal{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{rename\_tac}\hypertarget{method.rename-tac}{\hyperlink{method.rename-tac}{\mbox{\isa{rename{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{rotate\_tac}\hypertarget{method.rotate-tac}{\hyperlink{method.rotate-tac}{\mbox{\isa{rotate{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{tactic}\hypertarget{method.tactic}{\hyperlink{method.tactic}{\mbox{\isa{tactic}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{}{method}{raw\_tactic}\hypertarget{method.raw-tactic}{\hyperlink{method.raw-tactic}{\mbox{\isa{raw{\isaliteral{5F}{\isacharunderscore}}tactic}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{9}{} -\rail@bar -\rail@term{\hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.erule-tac}{\mbox{\isa{erule{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.drule-tac}{\mbox{\isa{drule{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{method.frule-tac}{\mbox{\isa{frule{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{method.cut-tac}{\mbox{\isa{cut{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{method.thin-tac}{\mbox{\isa{thin{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@cr{7} -\rail@bar -\rail@nont{\isa{dynamic{\isaliteral{5F}{\isacharunderscore}}insts}}[] -\rail@term{\isa{\isakeyword{in}}}[] -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@nextbar{8} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.subgoal-tac}{\mbox{\isa{subgoal{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.rename-tac}{\mbox{\isa{rename{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.rotate-tac}{\mbox{\isa{rotate{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.int}{\mbox{\isa{int}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{method.tactic}{\mbox{\isa{tactic}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.raw-tactic}{\mbox{\isa{raw{\isaliteral{5F}{\isacharunderscore}}tactic}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\rail@begin{2}{\isa{dynamic{\isaliteral{5F}{\isacharunderscore}}insts}} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - -\begin{description} - - \item \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}} etc. do resolution of rules with explicit - instantiation. This works the same way as the ML tactics \verb|res_inst_tac| etc. (see \cite{isabelle-implementation}) - - Multiple rules may be only given if there is no instantiation; then - \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}} is the same as \verb|resolve_tac| in ML (see - \cite{isabelle-implementation}). - - \item \hyperlink{method.cut-tac}{\mbox{\isa{cut{\isaliteral{5F}{\isacharunderscore}}tac}}} inserts facts into the proof state as - assumption of a subgoal; instantiations may be given as well. Note - that the scope of schematic variables is spread over the main goal - statement and rule premises are turned into new subgoals. This is - in contrast to the regular method \hyperlink{method.insert}{\mbox{\isa{insert}}} which inserts - closed rule statements. - - \item \hyperlink{method.thin-tac}{\mbox{\isa{thin{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} deletes the specified premise - from a subgoal. Note that \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} may contain schematic - variables, to abbreviate the intended proposition; the first - matching subgoal premise will be deleted. Removing useless premises - from a subgoal increases its readability and can make search tactics - run faster. - - \item \hyperlink{method.subgoal-tac}{\mbox{\isa{subgoal{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} adds the propositions - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} as local premises to a subgoal, and poses the same - as new subgoals (in the original context). - - \item \hyperlink{method.rename-tac}{\mbox{\isa{rename{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} renames parameters of a - goal according to the list \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, which refers to the - \emph{suffix} of variables. - - \item \hyperlink{method.rotate-tac}{\mbox{\isa{rotate{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{n} rotates the premises of a - subgoal by \isa{n} positions: from right to left if \isa{n} is - positive, and from left to right if \isa{n} is negative; the - default value is 1. - - \item \hyperlink{method.tactic}{\mbox{\isa{tactic}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} produces a proof method from - any ML text of type \verb|tactic|. Apart from the usual ML - environment and the current proof context, the ML code may refer to - the locally bound values \verb|facts|, which indicates any - current facts used for forward-chaining. - - \item \hyperlink{method.raw-tactic}{\mbox{\isa{raw{\isaliteral{5F}{\isacharunderscore}}tactic}}} is similar to \hyperlink{method.tactic}{\mbox{\isa{tactic}}}, but - presents the goal state in its raw internal form, where simultaneous - subgoals appear as conjunction of the logical framework instead of - the usual split into several subgoals. While feature this is useful - for debugging of complex method definitions, it should not never - appear in production theories. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{The Simplifier \label{sec:simplifier}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Simplification methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{simp}\hypertarget{method.simp}{\hyperlink{method.simp}{\mbox{\isa{simp}}}} & : & \isa{method} \\ - \indexdef{}{method}{simp\_all}\hypertarget{method.simp-all}{\hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{method.simp}{\mbox{\isa{simp}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{opt}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.simpmod}{\mbox{\isa{simpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{4}{\isa{opt}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm}}[] -\rail@nextbar{1} -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@nextbar{2} -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}use}}[] -\rail@nextbar{3} -\rail@term{\isa{asm{\isaliteral{5F}{\isacharunderscore}}lr}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{9}{\indexdef{}{syntax}{simpmod}\hypertarget{syntax.simpmod}{\hyperlink{syntax.simpmod}{\mbox{\isa{simpmod}}}}} -\rail@bar -\rail@term{\isa{add}}[] -\rail@nextbar{1} -\rail@term{\isa{del}}[] -\rail@nextbar{2} -\rail@term{\isa{only}}[] -\rail@nextbar{3} -\rail@term{\isa{cong}}[] -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{add}}[] -\rail@nextbar{5} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@nextbar{6} -\rail@term{\isa{split}}[] -\rail@bar -\rail@nextbar{7} -\rail@term{\isa{add}}[] -\rail@nextbar{8} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.simp}{\mbox{\isa{simp}}} invokes the Simplifier, after declaring - additional rules according to the arguments given. Note that the - \isa{only} modifier first removes all other rewrite rules, - congruences, and looper tactics (including splits), and then behaves - like \isa{add}. - - \medskip The \isa{cong} modifiers add or delete Simplifier - congruence rules (see also \secref{sec:simp-cong}), the default is - to add. - - \medskip The \isa{split} modifiers add or delete rules for the - Splitter (see also \cite{isabelle-ref}), the default is to add. - This works only if the Simplifier method has been properly setup to - include the Splitter (all major object logics such HOL, HOLCF, FOL, - ZF do this already). - - \item \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} is similar to \hyperlink{method.simp}{\mbox{\isa{simp}}}, but acts on - all goals (backwards from the last to the first one). - - \end{description} - - By default the Simplifier methods take local assumptions fully into - account, using equational assumptions in the subsequent - normalization process, or simplifying assumptions themselves (cf.\ - \verb|asm_full_simp_tac| in \cite{isabelle-ref}). In structured - proofs this is usually quite well behaved in practice: just the - local premises of the actual goal are involved, additional facts may - be inserted via explicit forward-chaining (via \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}, - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}, \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}} etc.). - - Additional Simplifier options may be specified to tune the behavior - further (mostly for unstructured scripts with many accidental local - facts): ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' means assumptions are ignored - completely (cf.\ \verb|simp_tac|), ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}simp{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' means - assumptions are used in the simplification of the conclusion but are - not themselves simplified (cf.\ \verb|asm_simp_tac|), and ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}use{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' means assumptions are simplified but are not used - in the simplification of each other or the conclusion (cf.\ \verb|full_simp_tac|). For compatibility reasons, there is also an option - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}asm{\isaliteral{5F}{\isacharunderscore}}lr{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'', which means that an assumption is only used - for simplifying assumptions which are to the right of it (cf.\ \verb|asm_lr_simp_tac|). - - The configuration option \isa{{\isaliteral{22}{\isachardoublequote}}depth{\isaliteral{5F}{\isacharunderscore}}limit{\isaliteral{22}{\isachardoublequote}}} limits the number of - recursive invocations of the simplifier during conditional - rewriting. - - \medskip The Splitter package is usually configured to work as part - of the Simplifier. The effect of repeatedly applying \verb|split_tac| can be simulated by ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}simp\ only{\isaliteral{3A}{\isacharcolon}}\ split{\isaliteral{3A}{\isacharcolon}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}''. There is also a separate \isa{split} - method available for single-step case splitting.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Declaring rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{print\_simpset}\hypertarget{command.print-simpset}{\hyperlink{command.print-simpset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}simpset}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{simp}\hypertarget{attribute.simp}{\hyperlink{attribute.simp}{\mbox{\isa{simp}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{split}\hypertarget{attribute.split}{\hyperlink{attribute.split}{\mbox{\isa{split}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{attribute.simp}{\mbox{\isa{simp}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.split}{\mbox{\isa{split}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.print-simpset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}simpset}}}} prints the collection of rules - declared to the Simplifier, which is also known as ``simpset'' - internally \cite{isabelle-ref}. - - \item \hyperlink{attribute.simp}{\mbox{\isa{simp}}} declares simplification rules. - - \item \hyperlink{attribute.split}{\mbox{\isa{split}}} declares case split rules. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Congruence rules\label{sec:simp-cong}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{attribute}{cong}\hypertarget{attribute.cong}{\hyperlink{attribute.cong}{\mbox{\isa{cong}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.cong}{\mbox{\isa{cong}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{attribute.cong}{\mbox{\isa{cong}}} declares congruence rules to the Simplifier - context. - - \end{description} - - Congruence rules are equalities of the form \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ f\ {\isaliteral{3F}{\isacharquery}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3F}{\isacharquery}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ f\ {\isaliteral{3F}{\isacharquery}}y\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3F}{\isacharquery}}y\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - This controls the simplification of the arguments of \isa{f}. For - example, some arguments can be simplified under additional - assumptions: \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Given this rule, the simplifier assumes \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} and extracts - rewrite rules from it when simplifying \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}. Such local - assumptions are effective for rewriting formulae such as \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ y\ {\isaliteral{2B}{\isacharplus}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}}. - - %FIXME - %The local assumptions are also provided as theorems to the solver; - %see \secref{sec:simp-solver} below. - - \medskip The following congruence rule for bounded quantifiers also - supplies contextual information --- about the bound variable: - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}A\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{3F}{\isacharquery}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}P\ x\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{3F}{\isacharquery}}Q\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{3F}{\isacharquery}}A{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{3F}{\isacharquery}}B{\isaliteral{2E}{\isachardot}}\ {\isaliteral{3F}{\isacharquery}}Q\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - \medskip This congruence rule for conditional expressions can - supply contextual information for simplifying the arms: - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}p\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}a\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{3F}{\isacharquery}}q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}b\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}d{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\isanewline -\ \ \ \ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{3F}{\isacharquery}}p\ then\ {\isaliteral{3F}{\isacharquery}}a\ else\ {\isaliteral{3F}{\isacharquery}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{3F}{\isacharquery}}q\ then\ {\isaliteral{3F}{\isacharquery}}c\ else\ {\isaliteral{3F}{\isacharquery}}d{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - A congruence rule can also \emph{prevent} simplification of some - arguments. Here is an alternative congruence rule for conditional - expressions that conforms to non-strict functional evaluation: - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}p\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{3F}{\isacharquery}}p\ then\ {\isaliteral{3F}{\isacharquery}}a\ else\ {\isaliteral{3F}{\isacharquery}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ {\isaliteral{3F}{\isacharquery}}q\ then\ {\isaliteral{3F}{\isacharquery}}a\ else\ {\isaliteral{3F}{\isacharquery}}b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Only the first argument is simplified; the others remain unchanged. - This can make simplification much faster, but may require an extra - case split over the condition \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}q{\isaliteral{22}{\isachardoublequote}}} to prove the goal.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Simplification procedures% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Simplification procedures are ML functions that produce proven - rewrite rules on demand. They are associated with higher-order - patterns that approximate the left-hand sides of equations. The - Simplifier first matches the current redex against one of the LHS - patterns; if this succeeds, the corresponding ML function is - invoked, passing the Simplifier context and redex term. Thus rules - may be specifically fashioned for particular situations, resulting - in a more powerful mechanism than term rewriting by a fixed set of - rules. - - Any successful result needs to be a (possibly conditional) rewrite - rule \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ u{\isaliteral{22}{\isachardoublequote}}} that is applicable to the current redex. The - rule will be applied just as any ordinary rewrite rule. It is - expected to be already in \emph{internal form}, bypassing the - automatic preprocessing of object-level equivalences. - - \begin{matharray}{rcl} - \indexdef{}{command}{simproc\_setup}\hypertarget{command.simproc-setup}{\hyperlink{command.simproc-setup}{\mbox{\isa{\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - simproc & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{6}{} -\rail@term{\hyperlink{command.simproc-setup}{\mbox{\isa{\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{identifier}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{5} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.simproc}{\mbox{\isa{simproc}}}}[] -\rail@bar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.simproc-setup}{\mbox{\isa{\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}}}} defines a named simplification - procedure that is invoked by the Simplifier whenever any of the - given term patterns match the current redex. The implementation, - which is provided as ML source text, needs to be of type \verb|morphism -> simpset -> cterm -> thm option|, where the \verb|cterm| represents the current redex \isa{r} and the result is - supposed to be some proven rewrite rule \isa{{\isaliteral{22}{\isachardoublequote}}r\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ r{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} (or a - generalized version), or \verb|NONE| to indicate failure. The - \verb|simpset| argument holds the full context of the current - Simplifier invocation, including the actual Isar proof context. The - \verb|morphism| informs about the difference of the original - compilation context wrt.\ the one of the actual application later - on. The optional \hyperlink{keyword.identifier}{\mbox{\isa{\isakeyword{identifier}}}} specifies theorems that - represent the logical content of the abstract theory of this - simproc. - - Morphisms and identifiers are only relevant for simprocs that are - defined within a local target context, e.g.\ in a locale. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}simproc\ add{\isaliteral{3A}{\isacharcolon}}\ name{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}simproc\ del{\isaliteral{3A}{\isacharcolon}}\ name{\isaliteral{22}{\isachardoublequote}}} - add or delete named simprocs to the current Simplifier context. The - default is to add a simproc. Note that \hyperlink{command.simproc-setup}{\mbox{\isa{\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}}}} - already adds the new simproc to the subsequent context. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Example% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following simplification procedure for \isa{{\isaliteral{28}{\isacharparenleft}}{\isaliteral{3F}{\isacharquery}}u{\isaliteral{5C3C436F6C6F6E3E}{\isasymColon}}unit{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}} in HOL performs fine-grained - control over rule application, beyond higher-order pattern matching. - Declaring \isa{unit{\isaliteral{5F}{\isacharunderscore}}eq} as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} directly would make - the simplifier loop! Note that a version of this simplification - procedure is already active in Isabelle/HOL.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -% -\endisadelimML -% -\isatagML -\isacommand{simproc{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ unit\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}x{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}unit{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ fn\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ ct\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ if\ HOLogic{\isaliteral{2E}{\isachardot}}is{\isaliteral{5F}{\isacharunderscore}}unit\ {\isaliteral{28}{\isacharparenleft}}term{\isaliteral{5F}{\isacharunderscore}}of\ ct{\isaliteral{29}{\isacharparenright}}\ then\ NONE\isanewline -\ \ \ \ else\ SOME\ {\isaliteral{28}{\isacharparenleft}}mk{\isaliteral{5F}{\isacharunderscore}}meta{\isaliteral{5F}{\isacharunderscore}}eq\ % -\isaantiq -thm\ unit{\isaliteral{5F}{\isacharunderscore}}eq{}% -\endisaantiq -{\isaliteral{29}{\isacharparenright}}\isanewline -{\isaliteral{2A7D}{\isacharverbatimclose}}% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\begin{isamarkuptext}% -Since the Simplifier applies simplification procedures - frequently, it is important to make the failure check in ML - reasonably fast.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Forward simplification% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{attribute}{simplified}\hypertarget{attribute.simplified}{\hyperlink{attribute.simplified}{\mbox{\isa{simplified}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.simplified}{\mbox{\isa{simplified}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{opt}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{opt}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm}}[] -\rail@nextbar{1} -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@nextbar{2} -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}use}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{attribute.simplified}{\mbox{\isa{simplified}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} causes a theorem to - be simplified, either by exactly the specified rules \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, or the implicit Simplifier context if no arguments are given. - The result is fully simplified by default, including assumptions and - conclusion; the options \isa{no{\isaliteral{5F}{\isacharunderscore}}asm} etc.\ tune the Simplifier in - the same way as the for the \isa{simp} method. - - Note that forward simplification restricts the simplifier to its - most basic operation of term rewriting; solver and looper tactics - \cite{isabelle-ref} are \emph{not} involved here. The \isa{simplified} attribute should be only rarely required under normal - circumstances. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{The Classical Reasoner \label{sec:classical}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Basic concepts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Although Isabelle is generic, many users will be working in - some extension of classical first-order logic. Isabelle/ZF is built - upon theory FOL, while Isabelle/HOL conceptually contains - first-order logic as a fragment. Theorem-proving in predicate logic - is undecidable, but many automated strategies have been developed to - assist in this task. - - Isabelle's classical reasoner is a generic package that accepts - certain information about a logic and delivers a suite of automatic - proof tools, based on rules that are classified and declared in the - context. These proof procedures are slow and simplistic compared - with high-end automated theorem provers, but they can save - considerable time and effort in practice. They can prove theorems - such as Pelletier's \cite{pelletier86} problems 40 and 41 in a few - milliseconds (including full proof reconstruction):% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ F\ x\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ F\ x\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ F\ z\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ F\ z\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}y{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ y\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ f\ x\ z\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ f\ x\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}z{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ z{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -The proof tools are generic. They are not restricted to - first-order logic, and have been heavily used in the development of - the Isabelle/HOL library and applications. The tactics can be - traced, and their components can be called directly; in this manner, - any proof can be viewed interactively.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{The sequent calculus% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle supports natural deduction, which is easy to use for - interactive proof. But natural deduction does not easily lend - itself to automation, and has a bias towards intuitionism. For - certain proofs in classical logic, it can not be called natural. - The \emph{sequent calculus}, a generalization of natural deduction, - is easier to automate. - - A \textbf{sequent} has the form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}, where \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}} - and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}} are sets of formulae.\footnote{For first-order - logic, sequents can equivalently be made from lists or multisets of - formulae.} The sequent \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} is - \textbf{valid} if \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} implies \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Thus \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} represent assumptions, each of which - is true, while \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} represent alternative goals. A - sequent is \textbf{basic} if its left and right sides have a common - formula, as in \isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ Q\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ R{\isaliteral{22}{\isachardoublequote}}}; basic sequents are trivially - valid. - - Sequent rules are classified as \textbf{right} or \textbf{left}, - indicating which side of the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}{\isaliteral{22}{\isachardoublequote}}} symbol they operate on. - Rules that operate on the right side are analogous to natural - deduction's introduction rules, and left rules are analogous to - elimination rules. The sequent calculus analogue of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - is the rule - \[ - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ Q{\isaliteral{22}{\isachardoublequote}}}} - \] - Applying the rule backwards, this breaks down some implication on - the right side of a sequent; \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}} stand for - the sets of formulae that are unaffected by the inference. The - analogue of the pair \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is the - single rule - \[ - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{2C}{\isacharcomma}}\ Q{\isaliteral{22}{\isachardoublequote}}}} - \] - This breaks down some disjunction on the right side, replacing it by - both disjuncts. Thus, the sequent calculus is a kind of - multiple-conclusion logic. - - To illustrate the use of multiple formulae on the right, let us - prove the classical theorem \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. Working - backwards, we reduce this formula to a basic sequent: - \[ - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} - {\infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} - {\infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}} - {\isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ Q\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{22}{\isachardoublequote}}}}}} - \] - - This example is typical of the sequent calculus: start with the - desired theorem and apply rules backwards in a fairly arbitrary - manner. This yields a surprisingly effective proof procedure. - Quantifiers add only few complications, since Isabelle handles - parameters and schematic variables. See \cite[Chapter - 10]{paulson-ml2} for further discussion.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Simulating sequents by natural deduction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle can represent sequents directly, as in the - object-logic LK. But natural deduction is easier to work with, and - most object-logics employ it. Fortunately, we can simulate the - sequent \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} by the Isabelle formula - \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} where the order of - the assumptions and the choice of \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} are arbitrary. - Elim-resolution plays a key role in simulating sequent proofs. - - We can easily handle reasoning on the left. Elim-resolution with - the rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} achieves - a similar effect as the corresponding sequent rules. For the other - connectives, we use sequent-style elimination rules instead of - destruction rules such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. - But note that the rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} has no effect under our - representation of sequents! - \[ - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ P{\isaliteral{22}{\isachardoublequote}}}} - \] - - What about reasoning on the right? Introduction rules can only - affect the formula in the conclusion, namely \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. The - other right-side formulae are represented as negated assumptions, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. In order to operate on one of these, it - must first be exchanged with \isa{{\isaliteral{22}{\isachardoublequote}}Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. Elim-resolution with the - \isa{swap} rule has this effect: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}} - - To ensure that swaps occur only when necessary, each introduction - rule is converted into a swapped form: it is resolved with the - second premise of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}swap{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. The swapped form of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, which might be called \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, is - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Similarly, the swapped form of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Swapped introduction rules are applied using elim-resolution, which - deletes the negated formula. Our representation of sequents also - requires the use of ordinary introduction rules. If we had no - regard for readability of intermediate goal states, we could treat - the right side more uniformly by representing sequents as \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F74746F6D3E}{\isasymbottom}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Extra rules for the sequent calculus% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As mentioned, destruction rules such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616E643E}{\isasymand}}E{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} must be replaced by sequent-style elimination rules. - In addition, we need rules to embody the classical equivalence - between \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}}. The introduction - rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}I{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} are replaced by a rule that simulates - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}: \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - The destruction rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is replaced by \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Quantifier replication also requires special rules. In classical - logic, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} is equivalent to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}; - the rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} are dual: - \[ - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ P\ t{\isaliteral{22}{\isachardoublequote}}}} - \qquad - \infer[\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}]{\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\ t{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C47616D6D613E}{\isasymGamma}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C44656C74613E}{\isasymDelta}}{\isaliteral{22}{\isachardoublequote}}}} - \] - Thus both kinds of quantifier may be replicated. Theorems requiring - multiple uses of a universal formula are easy to invent; consider - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}f\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ P\ a\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}f\isaliteral{5C3C5E7375703E}{}\isactrlsup n\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} for any - \isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3E}{\isachargreater}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. Natural examples of the multiple use of an - existential formula are rare; a standard one is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}y{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ y{\isaliteral{22}{\isachardoublequote}}}. - - Forgoing quantifier replication loses completeness, but gains - decidability, since the search space becomes finite. Many useful - theorems can be proved without replication, and the search generally - delivers its verdict in a reasonable time. To adopt this approach, - represent the sequent rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}L{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}R{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, - respectively, and put \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} into elimination form: \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - Elim-resolution with this rule will delete the universal formula - after a single use. To replicate universal quantifiers, replace the - rule by \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - To replicate existential quantifiers, replace \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}I{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} by - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - All introduction rules mentioned above are also useful in swapped - form. - - Replication makes the search space infinite; we must apply the rules - with care. The classical reasoner distinguishes between safe and - unsafe rules, applying the latter only when there is no alternative. - Depth-first search may well go down a blind alley; best-first search - is better behaved in an infinite search space. However, quantifier - replication is too expensive to prove any but the simplest theorems.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Rule declarations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The proof tools of the Classical Reasoner depend on - collections of rules declared in the context, which are classified - as introduction, elimination or destruction and as \emph{safe} or - \emph{unsafe}. In general, safe rules can be attempted blindly, - while unsafe rules must be used with care. A safe rule must never - reduce a provable goal to an unprovable set of subgoals. - - The rule \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} is unsafe because it reduces \isa{{\isaliteral{22}{\isachardoublequote}}P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} to \isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{22}{\isachardoublequote}}}, which might turn out as premature choice of an - unprovable subgoal. Any rule is unsafe whose premises contain new - unknowns. The elimination rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} is - unsafe, since it is applied via elim-resolution, which discards the - assumption \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} and replaces it by the weaker - assumption \isa{{\isaliteral{22}{\isachardoublequote}}P\ t{\isaliteral{22}{\isachardoublequote}}}. The rule \isa{{\isaliteral{22}{\isachardoublequote}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} is - unsafe for similar reasons. The quantifier duplication rule \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ t\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequote}}} is unsafe in a different sense: - since it keeps the assumption \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}, it is prone to - looping. In classical first-order logic, all rules are safe except - those mentioned above. - - The safe~/ unsafe distinction is vague, and may be regarded merely - as a way of giving some rules priority over others. One could argue - that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6F723E}{\isasymor}}E{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is unsafe, because repeated application of it - could generate exponentially many subgoals. Induction rules are - unsafe because inductive proofs are difficult to set up - automatically. Any inference is unsafe that instantiates an unknown - in the proof state --- thus matching must be used, rather than - unification. Even proof by assumption is unsafe if it instantiates - unknowns shared with other subgoals. - - \begin{matharray}{rcl} - \indexdef{}{command}{print\_claset}\hypertarget{command.print-claset}{\hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{intro}\hypertarget{attribute.intro}{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{elim}\hypertarget{attribute.elim}{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{dest}\hypertarget{attribute.dest}{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{rule}\hypertarget{attribute.rule}{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{iff}\hypertarget{attribute.iff}{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{swapped}\hypertarget{attribute.swapped}{\hyperlink{attribute.swapped}{\mbox{\isa{swapped}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{1} -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.rule}{\mbox{\isa{rule}}}}[] -\rail@term{\isa{del}}[] -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.iff}{\mbox{\isa{iff}}}}[] -\rail@bar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.print-claset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}claset}}}} prints the collection of rules - declared to the Classical Reasoner, i.e.\ the \verb|claset| - within the context. - - \item \hyperlink{attribute.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.elim}{\mbox{\isa{elim}}}, and \hyperlink{attribute.dest}{\mbox{\isa{dest}}} - declare introduction, elimination, and destruction rules, - respectively. By default, rules are considered as \emph{unsafe} - (i.e.\ not applied blindly without backtracking), while ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' classifies as \emph{safe}. Rule declarations marked by - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' coincide with those of Isabelle/Pure, cf.\ - \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps - of the \hyperlink{method.rule}{\mbox{\isa{rule}}} method). The optional natural number - specifies an explicit weight argument, which is ignored by the - automated reasoning tools, but determines the search order of single - rule steps. - - Introduction rules are those that can be applied using ordinary - resolution. Their swapped forms are generated internally, which - will be applied using elim-resolution. Elimination rules are - applied using elim-resolution. Rules are sorted by the number of - new subgoals they will yield; rules that generate the fewest - subgoals will be tried first. Otherwise, later declarations take - precedence over earlier ones. - - Rules already present in the context with the same classification - are ignored. A warning is printed if the rule has already been - added with some other classification, but the rule is added anyway - as requested. - - \item \hyperlink{attribute.rule}{\mbox{\isa{rule}}}~\isa{del} deletes all occurrences of a - rule from the classical context, regardless of its classification as - introduction~/ elimination~/ destruction and safe~/ unsafe. - - \item \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares logical equivalences to the - Simplifier and the Classical reasoner at the same time. - Non-conditional rules result in a safe introduction and elimination - pair; conditional ones are considered unsafe. Rules with negative - conclusion are automatically inverted (using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{22}{\isachardoublequote}}}-elimination - internally). - - The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' version of \hyperlink{attribute.iff}{\mbox{\isa{iff}}} declares rules to - the Isabelle/Pure context only, and omits the Simplifier - declaration. - - \item \hyperlink{attribute.swapped}{\mbox{\isa{swapped}}} turns an introduction rule into an - elimination, by resolving with the classical swap principle \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ R\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequote}}} in the second position. This is mainly for - illustrative purposes: the Classical Reasoner already swaps rules - internally as explained above. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Structured methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{rule}\hypertarget{method.rule}{\hyperlink{method.rule}{\mbox{\isa{rule}}}} & : & \isa{method} \\ - \indexdef{}{method}{contradiction}\hypertarget{method.contradiction}{\hyperlink{method.contradiction}{\mbox{\isa{contradiction}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.rule}{\mbox{\isa{rule}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.rule}{\mbox{\isa{rule}}} as offered by the Classical Reasoner is a - refinement over the Pure one (see \secref{sec:pure-meth-att}). Both - versions work the same, but the classical version observes the - classical rule context in addition to that of Isabelle/Pure. - - Common object logics (HOL, ZF, etc.) declare a rich collection of - classical rules (even if these would qualify as intuitionistic - ones), but only few declarations to the rule context of - Isabelle/Pure (\secref{sec:pure-meth-att}). - - \item \hyperlink{method.contradiction}{\mbox{\isa{contradiction}}} solves some goal by contradiction, - deriving any result from both \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequote}}} and \isa{A}. Chained - facts, which are guaranteed to participate, may appear in either - order. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Automated methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{blast}\hypertarget{method.blast}{\hyperlink{method.blast}{\mbox{\isa{blast}}}} & : & \isa{method} \\ - \indexdef{}{method}{auto}\hypertarget{method.auto}{\hyperlink{method.auto}{\mbox{\isa{auto}}}} & : & \isa{method} \\ - \indexdef{}{method}{force}\hypertarget{method.force}{\hyperlink{method.force}{\mbox{\isa{force}}}} & : & \isa{method} \\ - \indexdef{}{method}{fast}\hypertarget{method.fast}{\hyperlink{method.fast}{\mbox{\isa{fast}}}} & : & \isa{method} \\ - \indexdef{}{method}{slow}\hypertarget{method.slow}{\hyperlink{method.slow}{\mbox{\isa{slow}}}} & : & \isa{method} \\ - \indexdef{}{method}{best}\hypertarget{method.best}{\hyperlink{method.best}{\mbox{\isa{best}}}} & : & \isa{method} \\ - \indexdef{}{method}{fastforce}\hypertarget{method.fastforce}{\hyperlink{method.fastforce}{\mbox{\isa{fastforce}}}} & : & \isa{method} \\ - \indexdef{}{method}{slowsimp}\hypertarget{method.slowsimp}{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}} & : & \isa{method} \\ - \indexdef{}{method}{bestsimp}\hypertarget{method.bestsimp}{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}} & : & \isa{method} \\ - \indexdef{}{method}{deepen}\hypertarget{method.deepen}{\hyperlink{method.deepen}{\mbox{\isa{deepen}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.blast}{\mbox{\isa{blast}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.auto}{\mbox{\isa{auto}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.force}{\mbox{\isa{force}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{method.fast}{\mbox{\isa{fast}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.slow}{\mbox{\isa{slow}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.best}{\mbox{\isa{best}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{method.fastforce}{\mbox{\isa{fastforce}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.deepen}{\mbox{\isa{deepen}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] -\rail@endplus -\rail@end -\rail@begin{4}{\indexdef{}{syntax}{clamod}\hypertarget{syntax.clamod}{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}} -\rail@bar -\rail@bar -\rail@term{\isa{intro}}[] -\rail@nextbar{1} -\rail@term{\isa{elim}}[] -\rail@nextbar{2} -\rail@term{\isa{dest}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{1} -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextbar{3} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{14}{\indexdef{}{syntax}{clasimpmod}\hypertarget{syntax.clasimpmod}{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}} -\rail@bar -\rail@term{\isa{simp}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@nextbar{3} -\rail@term{\isa{only}}[] -\rail@endbar -\rail@nextbar{4} -\rail@bar -\rail@term{\isa{cong}}[] -\rail@nextbar{5} -\rail@term{\isa{split}}[] -\rail@endbar -\rail@bar -\rail@nextbar{5} -\rail@term{\isa{add}}[] -\rail@nextbar{6} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@nextbar{7} -\rail@term{\isa{iff}}[] -\rail@bar -\rail@bar -\rail@nextbar{8} -\rail@term{\isa{add}}[] -\rail@endbar -\rail@bar -\rail@nextbar{8} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextbar{9} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@nextbar{10} -\rail@bar -\rail@bar -\rail@term{\isa{intro}}[] -\rail@nextbar{11} -\rail@term{\isa{elim}}[] -\rail@nextbar{12} -\rail@term{\isa{dest}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{11} -\rail@nextbar{12} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@nextbar{13} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.blast}{\mbox{\isa{blast}}} is a separate classical tableau prover that - uses the same classical rule declarations as explained before. - - Proof search is coded directly in ML using special data structures. - A successful proof is then reconstructed using regular Isabelle - inferences. It is faster and more powerful than the other classical - reasoning tools, but has major limitations too. - - \begin{itemize} - - \item It does not use the classical wrapper tacticals, such as the - integration with the Simplifier of \hyperlink{method.fastforce}{\mbox{\isa{fastforce}}}. - - \item It does not perform higher-order unification, as needed by the - rule \isa{{\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ range\ {\isaliteral{3F}{\isacharquery}}f} in HOL. There are often - alternatives to such rules, for example \isa{{\isaliteral{3F}{\isacharquery}}b\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{3F}{\isacharquery}}b\ {\isaliteral{5C3C696E3E}{\isasymin}}\ range\ {\isaliteral{3F}{\isacharquery}}f}. - - \item Function variables may only be applied to parameters of the - subgoal. (This restriction arises because the prover does not use - higher-order unification.) If other function variables are present - then the prover will fail with the message \texttt{Function Var's - argument not a bound variable}. - - \item Its proof strategy is more general than \hyperlink{method.fast}{\mbox{\isa{fast}}} but can - be slower. If \hyperlink{method.blast}{\mbox{\isa{blast}}} fails or seems to be running forever, - try \hyperlink{method.fast}{\mbox{\isa{fast}}} and the other proof tools described below. - - \end{itemize} - - The optional integer argument specifies a bound for the number of - unsafe steps used in a proof. By default, \hyperlink{method.blast}{\mbox{\isa{blast}}} starts - with a bound of 0 and increases it successively to 20. In contrast, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}blast\ lim{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} tries to prove the goal using a search bound - of \isa{{\isaliteral{22}{\isachardoublequote}}lim{\isaliteral{22}{\isachardoublequote}}}. Sometimes a slow proof using \hyperlink{method.blast}{\mbox{\isa{blast}}} can - be made much faster by supplying the successful search bound to this - proof method instead. - - \item \hyperlink{method.auto}{\mbox{\isa{auto}}} combines classical reasoning with - simplification. It is intended for situations where there are a lot - of mostly trivial subgoals; it proves all the easy ones, leaving the - ones it cannot prove. Occasionally, attempting to prove the hard - ones may take a long time. - - The optional depth arguments in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}auto\ m\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} refer to its - builtin classical reasoning procedures: \isa{m} (default 4) is for - \hyperlink{method.blast}{\mbox{\isa{blast}}}, which is tried first, and \isa{n} (default 2) is - for a slower but more general alternative that also takes wrappers - into account. - - \item \hyperlink{method.force}{\mbox{\isa{force}}} is intended to prove the first subgoal - completely, using many fancy proof tools and performing a rather - exhaustive search. As a result, proof attempts may take rather long - or diverge easily. - - \item \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.best}{\mbox{\isa{best}}}, \hyperlink{method.slow}{\mbox{\isa{slow}}} attempt to - prove the first subgoal using sequent-style reasoning as explained - before. Unlike \hyperlink{method.blast}{\mbox{\isa{blast}}}, they construct proofs directly in - Isabelle. - - There is a difference in search strategy and back-tracking: \hyperlink{method.fast}{\mbox{\isa{fast}}} uses depth-first search and \hyperlink{method.best}{\mbox{\isa{best}}} uses best-first - search (guided by a heuristic function: normally the total size of - the proof state). - - Method \hyperlink{method.slow}{\mbox{\isa{slow}}} is like \hyperlink{method.fast}{\mbox{\isa{fast}}}, but conducts a broader - search: it may, when backtracking from a failed proof attempt, undo - even the step of proving a subgoal by assumption. - - \item \hyperlink{method.fastforce}{\mbox{\isa{fastforce}}}, \hyperlink{method.slowsimp}{\mbox{\isa{slowsimp}}}, \hyperlink{method.bestsimp}{\mbox{\isa{bestsimp}}} - are like \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.slow}{\mbox{\isa{slow}}}, \hyperlink{method.best}{\mbox{\isa{best}}}, - respectively, but use the Simplifier as additional wrapper. The name - \hyperlink{method.fastforce}{\mbox{\isa{fastforce}}}, reflects the behaviour of this popular method - better without requiring an understanding of its implementation. - - \item \hyperlink{method.deepen}{\mbox{\isa{deepen}}} works by exhaustive search up to a certain - depth. The start depth is 4 (unless specified explicitly), and the - depth is increased iteratively up to 10. Unsafe rules are modified - to preserve the formula they act on, so that it be used repeatedly. - This method can prove more goals than \hyperlink{method.fast}{\mbox{\isa{fast}}}, but is much - slower, for example if the assumptions have many universal - quantifiers. - - \end{description} - - Any of the above methods support additional modifiers of the context - of classical (and simplifier) rules, but the ones related to the - Simplifier are explicitly prefixed by \isa{simp} here. The - semantics of these ad-hoc rule declarations is analogous to the - attributes given before. Facts provided by forward chaining are - inserted into the goal before commencing proof search.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Semi-automated methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -These proof methods may help in situations when the - fully-automated tools fail. The result is a simpler subgoal that - can be tackled by other means, such as by manual instantiation of - quantifiers. - - \begin{matharray}{rcl} - \indexdef{}{method}{safe}\hypertarget{method.safe}{\hyperlink{method.safe}{\mbox{\isa{safe}}}} & : & \isa{method} \\ - \indexdef{}{method}{clarify}\hypertarget{method.clarify}{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}} & : & \isa{method} \\ - \indexdef{}{method}{clarsimp}\hypertarget{method.clarsimp}{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{method.safe}{\mbox{\isa{safe}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.clarify}{\mbox{\isa{clarify}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clamod}{\mbox{\isa{clamod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.safe}{\mbox{\isa{safe}}} repeatedly performs safe steps on all subgoals. - It is deterministic, with at most one outcome. - - \item \hyperlink{method.clarify}{\mbox{\isa{clarify}}} performs a series of safe steps without - splitting subgoals; see also \verb|clarify_step_tac|. - - \item \hyperlink{method.clarsimp}{\mbox{\isa{clarsimp}}} acts like \hyperlink{method.clarify}{\mbox{\isa{clarify}}}, but also does - simplification. Note that if the Simplifier context includes a - splitter for the premises, the subgoal may still be split. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Single-step tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{ML}{safe\_step\_tac}\verb|safe_step_tac: Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{inst\_step\_tac}\verb|inst_step_tac: Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{step\_tac}\verb|step_tac: Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{slow\_step\_tac}\verb|slow_step_tac: Proof.context -> int -> tactic| \\ - \indexdef{}{ML}{clarify\_step\_tac}\verb|clarify_step_tac: Proof.context -> int -> tactic| \\ - \end{matharray} - - These are the primitive tactics behind the (semi)automated proof - methods of the Classical Reasoner. By calling them yourself, you - can execute these procedures one step at a time. - - \begin{description} - - \item \verb|safe_step_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}ctxt\ i{\isaliteral{22}{\isachardoublequote}}} performs a safe step on - subgoal \isa{i}. The safe wrapper tacticals are applied to a - tactic that may include proof by assumption or Modus Ponens (taking - care not to instantiate unknowns), or substitution. - - \item \verb|inst_step_tac| is like \verb|safe_step_tac|, but allows - unknowns to be instantiated. - - \item \verb|step_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}ctxt\ i{\isaliteral{22}{\isachardoublequote}}} is the basic step of the proof - procedure. The unsafe wrapper tacticals are applied to a tactic - that tries \verb|safe_tac|, \verb|inst_step_tac|, or applies an unsafe - rule from the context. - - \item \verb|slow_step_tac| resembles \verb|step_tac|, but allows - backtracking between using safe rules with instantiation (\verb|inst_step_tac|) and using unsafe rules. The resulting search space - is larger. - - \item \verb|clarify_step_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}ctxt\ i{\isaliteral{22}{\isachardoublequote}}} performs a safe step - on subgoal \isa{i}. No splitting step is applied; for example, - the subgoal \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequote}}} is left as a conjunction. Proof by - assumption, Modus Ponens, etc., may be performed provided they do - not instantiate unknowns. Assumptions of the form \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} - may be eliminated. The safe wrapper tactical is applied. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Object-logic setup \label{sec:object-logic}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{judgment}\hypertarget{command.judgment}{\hyperlink{command.judgment}{\mbox{\isa{\isacommand{judgment}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{method}{atomize}\hypertarget{method.atomize}{\hyperlink{method.atomize}{\mbox{\isa{atomize}}}} & : & \isa{method} \\ - \indexdef{}{attribute}{atomize}\hypertarget{attribute.atomize}{\hyperlink{attribute.atomize}{\mbox{\isa{atomize}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{rule\_format}\hypertarget{attribute.rule-format}{\hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{rulify}\hypertarget{attribute.rulify}{\hyperlink{attribute.rulify}{\mbox{\isa{rulify}}}} & : & \isa{attribute} \\ - \end{matharray} - - The very starting point for any Isabelle object-logic is a ``truth - judgment'' that links object-level statements to the meta-logic - (with its minimal language of \isa{prop} that covers universal - quantification \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and implication \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}). - - Common object-logics are sufficiently expressive to internalize rule - statements over \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} within their own - language. This is useful in certain situations where a rule needs - to be viewed as an atomic statement from the meta-level perspective, - e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} versus \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequote}}}. - - From the following language elements, only the \hyperlink{method.atomize}{\mbox{\isa{atomize}}} - method and \hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}} attribute are occasionally - required by end-users, the rest is for those who need to setup their - own object-logic. In the latter case existing formulations of - Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. - - Generic tools may refer to the information provided by object-logic - declarations internally. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.judgment}{\mbox{\isa{\isacommand{judgment}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.atomize}{\mbox{\isa{atomize}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{full}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{noasm}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.judgment}{\mbox{\isa{\isacommand{judgment}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} declares constant - \isa{c} as the truth judgment of the current object-logic. Its - type \isa{{\isaliteral{5C3C7369676D613E}{\isasymsigma}}} should specify a coercion of the category of - object-level propositions to \isa{prop} of the Pure meta-logic; - the mixfix annotation \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} would typically just link the - object language (internally of syntactic category \isa{logic}) - with that of \isa{prop}. Only one \hyperlink{command.judgment}{\mbox{\isa{\isacommand{judgment}}}} - declaration may be given in any theory development. - - \item \hyperlink{method.atomize}{\mbox{\isa{atomize}}} (as a method) rewrites any non-atomic - premises of a sub-goal, using the meta-level equations declared via - \hyperlink{attribute.atomize}{\mbox{\isa{atomize}}} (as an attribute) beforehand. As a result, - heavily nested goals become amenable to fundamental operations such - as resolution (cf.\ the \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method). Giving the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}full{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option here means to turn the whole subgoal into an - object-statement (if possible), including the outermost parameters - and assumptions as well. - - A typical collection of \hyperlink{attribute.atomize}{\mbox{\isa{atomize}}} rules for a particular - object-logic would provide an internalization for each of the - connectives of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}, and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}}. - Meta-level conjunction should be covered as well (this is - particularly important for locales, see \secref{sec:locale}). - - \item \hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}} rewrites a theorem by the equalities - declared as \hyperlink{attribute.rulify}{\mbox{\isa{rulify}}} rules in the current object-logic. - By default, the result is fully normalized, including assumptions - and conclusions at any depth. The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option - restricts the transformation to the conclusion of a rule. - - In common object-logics (HOL, FOL, ZF), the effect of \hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}} is to replace (bounded) universal quantification - (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}}) and implication (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}) by the corresponding - rule statements over \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/HOLCF_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOLCF_Specific.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,129 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{HOLCF{\isaliteral{5F}{\isacharunderscore}}Specific}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ HOLCF{\isaliteral{5F}{\isacharunderscore}}Specific\isanewline -\isakeyword{imports}\ Base\ HOLCF\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Isabelle/HOLCF \label{ch:holcf}% -} -\isamarkuptrue% -% -\isamarkupsection{Mixfix syntax for continuous operations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOLCF}{command}{consts}\hypertarget{command.HOLCF.consts}{\hyperlink{command.HOLCF.consts}{\mbox{\isa{\isacommand{consts}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - HOLCF provides a separate type for continuous functions \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}}, with an explicit application operator \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C63646F743E}{\isasymcdot}}\ x{\isaliteral{22}{\isachardoublequote}}}. - Isabelle mixfix syntax normally refers directly to the pure - meta-level function type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}}, with application \isa{{\isaliteral{22}{\isachardoublequote}}f\ x{\isaliteral{22}{\isachardoublequote}}}. - - The HOLCF variant of \hyperlink{command.HOLCF.consts}{\mbox{\isa{\isacommand{consts}}}} modifies that of - Pure Isabelle (cf.\ \secref{sec:consts}) such that declarations - involving continuous function types are treated specifically. Any - given syntax template is transformed internally, generating - translation rules for the abstract and concrete representation of - continuous application. Note that mixing of HOLCF and Pure - application is \emph{not} supported!% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Recursive domains% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOLCF}{command}{domain}\hypertarget{command.HOLCF.domain}{\hyperlink{command.HOLCF.domain}{\mbox{\isa{\isacommand{domain}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOLCF.domain}{\mbox{\isa{\isacommand{domain}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\isa{spec}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{spec}} -\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@plus -\rail@nont{\isa{cons}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{cons}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\isa{dtrules}} -\rail@term{\isa{\isakeyword{distinct}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@term{\isa{\isakeyword{inject}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@term{\isa{\isakeyword{induction}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\end{railoutput} - - - Recursive domains in HOLCF are analogous to datatypes in classical - HOL (cf.\ \secref{sec:hol-datatype}). Mutual recursion is - supported, but no nesting nor arbitrary branching. Domain - constructors may be strict (default) or lazy, the latter admits to - introduce infinitary objects in the typical LCF manner (e.g.\ lazy - lists). See also \cite{MuellerNvOS99} for a general discussion of - HOLCF domains.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/HOL_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3512 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{HOL{\isaliteral{5F}{\isacharunderscore}}Specific}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ HOL{\isaliteral{5F}{\isacharunderscore}}Specific\isanewline -\isakeyword{imports}\ Base\ Main\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7E}{\isachartilde}}{\isaliteral{7E}{\isachartilde}}{\isaliteral{2F}{\isacharslash}}src{\isaliteral{2F}{\isacharslash}}HOL{\isaliteral{2F}{\isacharslash}}Library{\isaliteral{2F}{\isacharslash}}Old{\isaliteral{5F}{\isacharunderscore}}Recdef{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Isabelle/HOL \label{ch:hol}% -} -\isamarkuptrue% -% -\isamarkupsection{Higher-Order Logic% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/HOL is based on Higher-Order Logic, a polymorphic - version of Church's Simple Theory of Types. HOL can be best - understood as a simply-typed version of classical set theory. The - logic was first implemented in Gordon's HOL system - \cite{mgordon-hol}. It extends Church's original logic - \cite{church40} by explicit type variables (naive polymorphism) and - a sound axiomatization scheme for new types based on subsets of - existing types. - - Andrews's book \cite{andrews86} is a full description of the - original Church-style higher-order logic, with proofs of correctness - and completeness wrt.\ certain set-theoretic interpretations. The - particular extensions of Gordon-style HOL are explained semantically - in two chapters of the 1993 HOL book \cite{pitts93}. - - Experience with HOL over decades has demonstrated that higher-order - logic is widely applicable in many areas of mathematics and computer - science. In a sense, Higher-Order Logic is simpler than First-Order - Logic, because there are fewer restrictions and special cases. Note - that HOL is \emph{weaker} than FOL with axioms for ZF set theory, - which is traditionally considered the standard foundation of regular - mathematics, but for most applications this does not matter. If you - prefer ML to Lisp, you will probably prefer HOL to ZF. - - \medskip The syntax of HOL follows \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus and - functional programming. Function application is curried. To apply - the function \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}} to the - arguments \isa{a} and \isa{b} in HOL, you simply write \isa{{\isaliteral{22}{\isachardoublequote}}f\ a\ b{\isaliteral{22}{\isachardoublequote}}} (as in ML or Haskell). There is no ``apply'' operator; the - existing application of the Pure \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-calculus is re-used. - Note that in HOL \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} means ``\isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to - the pair \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} (which is notation for \isa{{\isaliteral{22}{\isachardoublequote}}Pair\ a\ b{\isaliteral{22}{\isachardoublequote}}}). The latter typically introduces extra formal efforts that can - be avoided by currying functions by default. Explicit tuples are as - infrequent in HOL formalizations as in good ML or Haskell programs. - - \medskip Isabelle/HOL has a distinct feel, compared to other - object-logics like Isabelle/ZF. It identifies object-level types - with meta-level types, taking advantage of the default - type-inference mechanism of Isabelle/Pure. HOL fully identifies - object-level functions with meta-level functions, with native - abstraction and application. - - These identifications allow Isabelle to support HOL particularly - nicely, but they also mean that HOL requires some sophistication - from the user. In particular, an understanding of Hindley-Milner - type-inference with type-classes, which are both used extensively in - the standard libraries and applications. Beginners can set - \hyperlink{attribute.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}} or even \hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}} to get more - explicit information about the result of type-inference.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Inductive and coinductive definitions \label{sec:hol-inductive}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{inductive}\hypertarget{command.HOL.inductive}{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{inductive\_set}\hypertarget{command.HOL.inductive-set}{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{coinductive}\hypertarget{command.HOL.coinductive}{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{coinductive\_set}\hypertarget{command.HOL.coinductive-set}{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{mono}\hypertarget{attribute.HOL.mono}{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}} & : & \isa{attribute} \\ - \end{matharray} - - An \emph{inductive definition} specifies the least predicate or set - \isa{R} closed under given rules: applying a rule to elements of - \isa{R} yields a result within \isa{R}. For example, a - structural operational semantics is an inductive definition of an - evaluation relation. - - Dually, a \emph{coinductive definition} specifies the greatest - predicate or set \isa{R} that is consistent with given rules: - every element of \isa{R} can be seen as arising by applying a rule - to elements of \isa{R}. An important example is using - bisimulation relations to formalise equivalence of processes and - infinite data structures. - - Both inductive and coinductive definitions are based on the - Knaster-Tarski fixed-point theorem for complete lattices. The - collection of introduction rules given by the user determines a - functor on subsets of set-theoretic relations. The required - monotonicity of the recursion scheme is proven as a prerequisite to - the fixed-point definition and the resulting consequences. This - works by pushing inclusion through logical connectives and any other - operator that might be wrapped around recursive occurrences of the - defined relation: there must be a monotonicity theorem of the form - \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for each premise \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}} in an - introduction rule. The default rule declarations of Isabelle/HOL - already take care of most common situations. - - \begin{railoutput} -\rail@begin{10}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@cr{5} -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@bar -\rail@nextbar{6} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{6} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{clauses}}[] -\rail@endbar -\rail@cr{8} -\rail@bar -\rail@nextbar{9} -\rail@term{\isa{\isakeyword{monos}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{clauses}} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.inductive}{\mbox{\isa{\isacommand{inductive}}}} and \hyperlink{command.HOL.coinductive}{\mbox{\isa{\isacommand{coinductive}}}} define (co)inductive predicates from the introduction - rules. - - The propositions given as \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}} in the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} part are either rules of the usual \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} format - (with arbitrary nesting), or equalities using \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}}. The - latter specifies extra-logical abbreviations in the sense of - \indexref{}{command}{abbreviation}\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}. Introducing abstract syntax - simultaneously with the actual introduction rules is occasionally - useful for complex specifications. - - The optional \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} part contains a list of parameters of - the (co)inductive predicates that remain fixed throughout the - definition, in contrast to arguments of the relation that may vary - in each occurrence within the given \isa{{\isaliteral{22}{\isachardoublequote}}clauses{\isaliteral{22}{\isachardoublequote}}}. - - The optional \hyperlink{keyword.monos}{\mbox{\isa{\isakeyword{monos}}}} declaration contains additional - \emph{monotonicity theorems}, which are required for each operator - applied to a recursive set in the introduction rules. - - \item \hyperlink{command.HOL.inductive-set}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}}}} and \hyperlink{command.HOL.coinductive-set}{\mbox{\isa{\isacommand{coinductive{\isaliteral{5F}{\isacharunderscore}}set}}}} are wrappers for to the previous commands for - native HOL predicates. This allows to define (co)inductive sets, - where multiple arguments are simulated via tuples. - - \item \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} declares monotonicity rules in the - context. These rule are involved in the automated monotonicity - proof of the above inductive and coinductive definitions. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A (co)inductive definition of \isa{R} provides the following - main theorems: - - \begin{description} - - \item \isa{R{\isaliteral{2E}{\isachardot}}intros} is the list of introduction rules as proven - theorems, for the recursive predicates (or sets). The rules are - also available individually, using the names given them in the - theory file; - - \item \isa{R{\isaliteral{2E}{\isachardot}}cases} is the case analysis (or elimination) rule; - - \item \isa{R{\isaliteral{2E}{\isachardot}}induct} or \isa{R{\isaliteral{2E}{\isachardot}}coinduct} is the (co)induction - rule; - - \item \isa{R{\isaliteral{2E}{\isachardot}}simps} is the equation unrolling the fixpoint of the - predicate one step. - - \end{description} - - When several predicates \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are - defined simultaneously, the list of introduction rules is called - \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}intros{\isaliteral{22}{\isachardoublequote}}}, the case analysis rules are - called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}cases{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}cases{\isaliteral{22}{\isachardoublequote}}}, and the list - of mutual induction rules is called \isa{{\isaliteral{22}{\isachardoublequote}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5F}{\isacharunderscore}}R\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}inducts{\isaliteral{22}{\isachardoublequote}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Monotonicity theorems% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The context maintains a default set of theorems that are used - in monotonicity proofs. New rules can be declared via the - \hyperlink{attribute.HOL.mono}{\mbox{\isa{mono}}} attribute. See the main Isabelle/HOL - sources for some examples. The general format of such monotonicity - theorems is as follows: - - \begin{itemize} - - \item Theorems of the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ A\ {\isaliteral{5C3C6C653E}{\isasymle}}\ {\isaliteral{5C3C4D3E}{\isasymM}}\ B{\isaliteral{22}{\isachardoublequote}}}, for proving - monotonicity of inductive definitions whose introduction rules have - premises involving terms such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4D3E}{\isasymM}}\ R\ t{\isaliteral{22}{\isachardoublequote}}}. - - \item Monotonicity theorems for logical operators, which are of the - general form \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}. For example, in - the case of the operator \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, the corresponding theorem is - \[ - \infer{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}}{\isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}P\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}} - \] - - \item De Morgan style equations for reasoning about the ``polarity'' - of expressions, e.g. - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ P{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ Q{\isaliteral{22}{\isachardoublequote}}} - \] - - \item Equations for reducing complex operators to more primitive - ones whose monotonicity can easily be proved, e.g. - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ Q{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6C6F6E676C65667472696768746172726F773E}{\isasymlongleftrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequote}}} \qquad\qquad - \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ P\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequote}}} - \] - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The finite powerset operator can be defined inductively like this:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}set}\isamarkupfalse% -\ Fin\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ set\ set{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{for}\ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ set{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\isanewline -\ \ empty{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{7D}{\isacharbraceright}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ insert{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ insert\ a\ B\ {\isaliteral{5C3C696E3E}{\isasymin}}\ Fin\ A{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -The accessible part of a relation is defined as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{inductive}\isamarkupfalse% -\ acc\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{for}\ r\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{28}{\isacharparenleft}}\isakeyword{infix}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C707265633E}{\isasymprec}}{\isaliteral{22}{\isachardoublequoteclose}}\ {\isadigit{5}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}\isanewline -\isakeyword{where}\ acc{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}y{\isaliteral{2E}{\isachardot}}\ y\ {\isaliteral{5C3C707265633E}{\isasymprec}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ acc\ r\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Common logical connectives can be easily characterized as -non-recursive inductive definitions with parameters, but without -arguments.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{inductive}\isamarkupfalse% -\ AND\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline -\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ AND\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{inductive}\isamarkupfalse% -\ OR\ \isakeyword{for}\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\isanewline -\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ OR\ A\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{inductive}\isamarkupfalse% -\ EXISTS\ \isakeyword{for}\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ a\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ EXISTS\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Here the \isa{{\isaliteral{22}{\isachardoublequote}}cases{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}induct{\isaliteral{22}{\isachardoublequote}}} rules produced by - the \hyperlink{command.inductive}{\mbox{\isa{\isacommand{inductive}}}} package coincide with the expected - elimination rules for Natural Deduction. Already in the original - article by Gerhard Gentzen \cite{Gentzen:1935} there is a hint that - each connective can be characterized by its introductions, and the - elimination can be constructed systematically.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Recursive functions \label{sec:recursion}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{primrec}\hypertarget{command.HOL.primrec}{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{fun}\hypertarget{command.HOL.fun}{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{function}\hypertarget{command.HOL.function}{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{termination}\hypertarget{command.HOL.termination}{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{equations}}[] -\rail@end -\rail@begin{4}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{functionopts}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@cr{3} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{equations}}[] -\rail@end -\rail@begin{3}{\isa{equations}} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{functionopts}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@bar -\rail@term{\isa{sequential}}[] -\rail@nextbar{1} -\rail@term{\isa{domintros}}[] -\rail@endbar -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}} defines primitive recursive - functions over datatypes (see also \indexref{HOL}{command}{datatype}\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} and - \indexref{HOL}{command}{rep\_datatype}\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}). The given \isa{equations} - specify reduction rules that are produced by instantiating the - generic combinator for primitive recursion that is available for - each datatype. - - Each equation needs to be of the form: - - \begin{isabelle}% -{\isaliteral{22}{\isachardoublequote}}f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{28}{\isacharparenleft}}C\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{29}{\isacharparenright}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ z\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ rhs{\isaliteral{22}{\isachardoublequote}}% -\end{isabelle} - - such that \isa{C} is a datatype constructor, \isa{rhs} contains - only the free variables on the left-hand side (or from the context), - and all recursive occurrences of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} in \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}} are of - the form \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} for some \isa{i}. At most one - reduction rule for each constructor can be given. The order does - not matter. For missing constructors, the function is defined to - return a default value, but this equation is made difficult to - access for users. - - The reduction rules are declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} by default, - which enables standard proof methods like \hyperlink{method.simp}{\mbox{\isa{simp}}} and - \hyperlink{method.auto}{\mbox{\isa{auto}}} to normalize expressions of \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} applied to - datatype constructions, by simulating symbolic computation via - rewriting. - - \item \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} defines functions by general - wellfounded recursion. A detailed description with examples can be - found in \cite{isabelle-function}. The function is specified by a - set of (possibly conditional) recursive equations with arbitrary - pattern matching. The command generates proof obligations for the - completeness and the compatibility of patterns. - - The defined function is considered partial, and the resulting - simplification rules (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}psimps{\isaliteral{22}{\isachardoublequote}}}) and induction rule - (named \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}pinduct{\isaliteral{22}{\isachardoublequote}}}) are guarded by a generated domain - predicate \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5F}{\isacharunderscore}}dom{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}} - command can then be used to establish that the function is total. - - \item \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} is a shorthand notation for ``\hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}sequential{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, followed by automated - proof attempts regarding pattern matching and termination. See - \cite{isabelle-function} for further details. - - \item \hyperlink{command.HOL.termination}{\mbox{\isa{\isacommand{termination}}}}~\isa{f} commences a - termination proof for the previously defined function \isa{f}. If - this is omitted, the command refers to the most recent function - definition. After the proof is closed, the recursive equations and - the induction principle is established. - - \end{description} - - Recursive definitions introduced by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} - command accommodate reasoning by induction (cf.\ \hyperlink{method.induct}{\mbox{\isa{induct}}}): - rule \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}induct{\isaliteral{22}{\isachardoublequote}}} refers to a specific induction rule, with - parameters named according to the user-specified equations. Cases - are numbered starting from 1. For \hyperlink{command.HOL.primrec}{\mbox{\isa{\isacommand{primrec}}}}, the - induction principle coincides with structural recursion on the - datatype where the recursion is carried out. - - The equations provided by these packages may be referred later as - theorem list \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, where \isa{f} is the (collective) - name of the functions defined. Individual equations may be named - explicitly as well. - - The \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} command accepts the following - options. - - \begin{description} - - \item \isa{sequential} enables a preprocessor which disambiguates - overlapping patterns by making them mutually disjoint. Earlier - equations take precedence over later ones. This allows to give the - specification in a format very similar to functional programming. - Note that the resulting simplification and induction rules - correspond to the transformed specification, not the one given - originally. This usually means that each equation given by the user - may result in several theorems. Also note that this automatic - transformation only works for ML-style datatype patterns. - - \item \isa{domintros} enables the automated generation of - introduction rules for the domain predicate. While mostly not - needed, they can be helpful in some proofs about partial functions. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Example: evaluation of expressions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Subsequently, we define mutual datatypes for arithmetic and - boolean expressions, and use \hyperlink{command.primrec}{\mbox{\isa{\isacommand{primrec}}}} for evaluation - functions that follow the same recursive structure.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ IF\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ Sum\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ Diff\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ Var\ {\isaliteral{27}{\isacharprime}}a\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ Num\ nat\isanewline -\isakeyword{and}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{3D}{\isacharequal}}\isanewline -\ \ \ \ Less\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ And\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ Neg\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -\medskip Evaluation of arithmetic and boolean expressions% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\ evala\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{and}\ evalb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}if\ evalb\ env\ b\ then\ evala\ env\ a{\isadigit{1}}\ else\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2B}{\isacharplus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isadigit{1}}\ {\isaliteral{2D}{\isacharminus}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ env\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evala\ env\ a{\isadigit{1}}\ {\isaliteral{3C}{\isacharless}}\ evala\ env\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}evalb\ env\ b{\isadigit{1}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ evalb\ env\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ evalb\ env\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Since the value of an expression depends on the value of its - variables, the functions \isa{evala} and \isa{evalb} take an - additional parameter, an \emph{environment} that maps variables to - their values. - - \medskip Substitution on expressions can be defined similarly. The - mapping \isa{f} of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} given as a - parameter is lifted canonically on the types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, respectively.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\ substa\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ aexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isakeyword{and}\ substb\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ aexp{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ bexp\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ bexp{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}IF\ b\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ IF\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Sum\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Sum\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Diff\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Diff\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ v{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ v{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substa\ f\ {\isaliteral{28}{\isacharparenleft}}Num\ n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Num\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Less\ a{\isadigit{1}}\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Less\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substa\ f\ a{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}And\ b{\isadigit{1}}\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ And\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}substb\ f\ {\isaliteral{28}{\isacharparenleft}}Neg\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Neg\ {\isaliteral{28}{\isacharparenleft}}substb\ f\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -In textbooks about semantics one often finds substitution - theorems, which express the relationship between substitution and - evaluation. For \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ aexp{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ bexp{\isaliteral{22}{\isachardoublequote}}}, we can prove - such a theorem by mutual induction, followed by simplification.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ subst{\isaliteral{5F}{\isacharunderscore}}one{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ {\isaliteral{28}{\isacharparenleft}}Var\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}env\ {\isaliteral{28}{\isacharparenleft}}v\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ evala\ env\ a{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ subst{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}evala\ env\ {\isaliteral{28}{\isacharparenleft}}substa\ s\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evala\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}evalb\ env\ {\isaliteral{28}{\isacharparenleft}}substb\ s\ b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ evalb\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ evala\ env\ {\isaliteral{28}{\isacharparenleft}}s\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ a\ \isakeyword{and}\ b{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsubsection{Example: a substitution function for terms% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Functions on datatypes with nested recursion are also defined - by mutual primitive recursion.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{22}{\isachardoublequoteopen}}term{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{3D}{\isacharequal}}\ Var\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ App\ {\isaliteral{27}{\isacharprime}}b\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -A substitution function on type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}} can be - defined as follows, by working simultaneously on \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequote}}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{primrec}\isamarkupfalse% -\ subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term\ list{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}Var\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ f\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ {\isaliteral{28}{\isacharparenleft}}App\ b\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ App\ b\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ {\isaliteral{28}{\isacharparenleft}}t\ {\isaliteral{23}{\isacharhash}}\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f\ t\ {\isaliteral{23}{\isacharhash}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f\ ts{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -The recursion scheme follows the structure of the unfolded - definition of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ term{\isaliteral{22}{\isachardoublequote}}}. To prove properties of this - substitution function, mutual induction is needed:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{2}}\ t{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term\ f{\isadigit{1}}\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ ts\ {\isaliteral{3D}{\isacharequal}}\ subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{1}}\ {\isaliteral{28}{\isacharparenleft}}subst{\isaliteral{5F}{\isacharunderscore}}term{\isaliteral{5F}{\isacharunderscore}}list\ f{\isadigit{2}}\ ts{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ t\ \isakeyword{and}\ ts{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsubsection{Example: a map function for infinitely branching trees% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Defining functions on infinitely branching datatypes by - primitive recursion is just as easy.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{7C}{\isacharbar}}\ Branch\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{primrec}\isamarkupfalse% -\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}a\ tree\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{27}{\isacharprime}}b\ tree{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isakeyword{where}\isanewline -\ \ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Atom\ a{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Atom\ {\isaliteral{28}{\isacharparenleft}}f\ a{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -{\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}Branch\ ts{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ Branch\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{28}{\isacharparenleft}}ts\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -Note that all occurrences of functions such as \isa{ts} - above must be applied to an argument. In particular, \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ ts{\isaliteral{22}{\isachardoublequote}}} is not allowed here.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\begin{isamarkuptext}% -Here is a simple composition lemma for \isa{map{\isaliteral{5F}{\isacharunderscore}}tree}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}map{\isaliteral{5F}{\isacharunderscore}}tree\ g\ {\isaliteral{28}{\isacharparenleft}}map{\isaliteral{5F}{\isacharunderscore}}tree\ f\ t{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ map{\isaliteral{5F}{\isacharunderscore}}tree\ {\isaliteral{28}{\isacharparenleft}}g\ {\isaliteral{5C3C636972633E}{\isasymcirc}}\ f{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ t{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Proof methods related to recursive definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{pat\_completeness}\hypertarget{method.HOL.pat-completeness}{\hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{relation}\hypertarget{method.HOL.relation}{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{lexicographic\_order}\hypertarget{method.HOL.lexicographic-order}{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{size\_change}\hypertarget{method.HOL.size-change}{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{induction\_schema}\hypertarget{method.HOL.induction-schema}{\hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}}}[] -\rail@nont{\isa{orders}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endplus -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}}}[] -\rail@end -\rail@begin{4}{\isa{orders}} -\rail@plus -\rail@nextplus{1} -\rail@bar -\rail@term{\isa{max}}[] -\rail@nextbar{2} -\rail@term{\isa{min}}[] -\rail@nextbar{3} -\rail@term{\isa{ms}}[] -\rail@endbar -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.pat-completeness}{\mbox{\isa{pat{\isaliteral{5F}{\isacharunderscore}}completeness}}} is a specialized method to - solve goals regarding the completeness of pattern matching, as - required by the \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} package (cf.\ - \cite{isabelle-function}). - - \item \hyperlink{method.HOL.relation}{\mbox{\isa{relation}}}~\isa{R} introduces a termination - proof using the relation \isa{R}. The resulting proof state will - contain goals expressing that \isa{R} is wellfounded, and that the - arguments of recursive calls decrease with respect to \isa{R}. - Usually, this method is used as the initial proof step of manual - termination proofs. - - \item \hyperlink{method.HOL.lexicographic-order}{\mbox{\isa{lexicographic{\isaliteral{5F}{\isacharunderscore}}order}}} attempts a fully - automated termination proof by searching for a lexicographic - combination of size measures on the arguments of the function. The - method accepts the same arguments as the \hyperlink{method.auto}{\mbox{\isa{auto}}} method, - which it uses internally to prove local descents. The \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}). - - In case of failure, extensive information is printed, which can help - to analyse the situation (cf.\ \cite{isabelle-function}). - - \item \hyperlink{method.HOL.size-change}{\mbox{\isa{size{\isaliteral{5F}{\isacharunderscore}}change}}} also works on termination goals, - using a variation of the size-change principle, together with a - graph decomposition technique (see \cite{krauss_phd} for details). - Three kinds of orders are used internally: \isa{max}, \isa{min}, - and \isa{ms} (multiset), which is only available when the theory - \isa{Multiset} is loaded. When no order kinds are given, they are - tried in order. The search for a termination proof uses SAT solving - internally. - - For local descent proofs, the \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} modifiers are - accepted (as for \hyperlink{method.auto}{\mbox{\isa{auto}}}). - - \item \hyperlink{method.HOL.induction-schema}{\mbox{\isa{induction{\isaliteral{5F}{\isacharunderscore}}schema}}} derives user-specified - induction rules from well-founded induction and completeness of - patterns. This factors out some operations that are done internally - by the function package and makes them available separately. See - \verb|~~/src/HOL/ex/Induction_Schema.thy| for examples. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Functions with explicit partiality% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{partial\_function}\hypertarget{command.HOL.partial-function}{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{partial\_function\_mono}\hypertarget{attribute.HOL.partial-function-mono}{\hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{5}{} -\rail@term{\hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@cr{3} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.partial-function}{\mbox{\isa{\isacommand{partial{\isaliteral{5F}{\isacharunderscore}}function}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} defines - recursive functions based on fixpoints in complete partial - orders. No termination proof is required from the user or - constructed internally. Instead, the possibility of non-termination - is modelled explicitly in the result type, which contains an - explicit bottom element. - - Pattern matching and mutual recursion are currently not supported. - Thus, the specification consists of a single function described by a - single recursive equation. - - There are no fixed syntactic restrictions on the body of the - function, but the induced functional must be provably monotonic - wrt.\ the underlying order. The monotonicitity proof is performed - internally, and the definition is rejected when it fails. The proof - can be influenced by declaring hints using the - \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} attribute. - - The mandatory \isa{mode} argument specifies the mode of operation - of the command, which directly corresponds to a complete partial - order on the result type. By default, the following modes are - defined: - - \begin{description} - - \item \isa{option} defines functions that map into the \isa{option} type. Here, the value \isa{None} is used to model a - non-terminating computation. Monotonicity requires that if \isa{None} is returned by a recursive call, then the overall result must - also be \isa{None}. This is best achieved through the use of the - monadic operator \isa{{\isaliteral{22}{\isachardoublequote}}Option{\isaliteral{2E}{\isachardot}}bind{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{tailrec} defines functions with an arbitrary result - type and uses the slightly degenerated partial order where \isa{{\isaliteral{22}{\isachardoublequote}}undefined{\isaliteral{22}{\isachardoublequote}}} is the bottom element. Now, monotonicity requires that - if \isa{undefined} is returned by a recursive call, then the - overall result must also be \isa{undefined}. In practice, this is - only satisfied when each recursive call is a tail call, whose result - is directly returned. Thus, this mode of operation allows the - definition of arbitrary tail-recursive functions. - - \end{description} - - Experienced users may define new modes by instantiating the locale - \isa{{\isaliteral{22}{\isachardoublequote}}partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}definitions{\isaliteral{22}{\isachardoublequote}}} appropriately. - - \item \hyperlink{attribute.HOL.partial-function-mono}{\mbox{\isa{partial{\isaliteral{5F}{\isacharunderscore}}function{\isaliteral{5F}{\isacharunderscore}}mono}}} declares rules for - use in the internal monononicity proofs of partial function - definitions. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Old-style recursive function definitions (TFL)% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{recdef}\hypertarget{command.HOL.recdef}{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{recdef\_tc}\hypertarget{command.HOL.recdef-tc}{\hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - The old TFL commands \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} and \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}} for defining recursive are mostly obsolete; \hyperlink{command.HOL.function}{\mbox{\isa{\isacommand{function}}}} or \hyperlink{command.HOL.fun}{\mbox{\isa{\isacommand{fun}}}} should be used instead. - - \begin{railoutput} -\rail@begin{5}{} -\rail@term{\hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{permissive}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{hints}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@nont{\isa{recdeftc}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\isa{tc}}[] -\rail@end -\rail@begin{2}{\isa{hints}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{hints}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{recdefmod}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{4}{\isa{recdefmod}} -\rail@bar -\rail@bar -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@nextbar{1} -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}[] -\rail@nextbar{2} -\rail@term{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{tc}} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} defines general well-founded - recursive functions (using the TFL package), see also - \cite{isabelle-HOL}. The ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}permissive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option tells - TFL to recover from failed proof attempts, returning unfinished - results. The \isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}, \isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}, and \isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf} hints refer to auxiliary rules to be used in the internal - automated proof process of TFL. Additional \hyperlink{syntax.clasimpmod}{\mbox{\isa{clasimpmod}}} - declarations may be given to tune the context of the Simplifier - (cf.\ \secref{sec:simplifier}) and Classical reasoner (cf.\ - \secref{sec:classical}). - - \item \hyperlink{command.HOL.recdef-tc}{\mbox{\isa{\isacommand{recdef{\isaliteral{5F}{\isacharunderscore}}tc}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} recommences the - proof for leftover termination condition number \isa{i} (default - 1) as generated by a \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} definition of - constant \isa{c}. - - Note that in most cases, \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} is able to finish - its internal proofs without manual intervention. - - \end{description} - - \medskip Hints for \hyperlink{command.HOL.recdef}{\mbox{\isa{\isacommand{recdef}}}} may be also declared - globally, using the following attributes. - - \begin{matharray}{rcl} - \indexdef{HOL}{attribute}{recdef\_simp}\hypertarget{attribute.HOL.recdef-simp}{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{recdef\_cong}\hypertarget{attribute.HOL.recdef-cong}{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{recdef\_wf}\hypertarget{attribute.HOL.recdef-wf}{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{attribute.HOL.recdef-simp}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}simp}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.HOL.recdef-cong}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}cong}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{attribute.HOL.recdef-wf}{\mbox{\isa{recdef{\isaliteral{5F}{\isacharunderscore}}wf}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Datatypes \label{sec:hol-datatype}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{datatype}\hypertarget{command.HOL.datatype}{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{rep\_datatype}\hypertarget{command.HOL.rep-datatype}{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}}}[] -\rail@plus -\rail@nont{\isa{spec}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{2} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{\isa{spec}} -\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@plus -\rail@nont{\isa{cons}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{cons}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} defines inductive datatypes in - HOL. - - \item \hyperlink{command.HOL.rep-datatype}{\mbox{\isa{\isacommand{rep{\isaliteral{5F}{\isacharunderscore}}datatype}}}} represents existing types as - datatypes. - - For foundational reasons, some basic types such as \isa{nat}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}, \isa{bool} and \isa{unit} are - introduced by more primitive means using \indexref{}{command}{typedef}\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}. To - recover the rich infrastructure of \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}} (e.g.\ rules - for \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} and the primitive recursion - combinators), such types may be represented as actual datatypes - later. This is done by specifying the constructors of the desired - type, and giving a proof of the induction rule, distinctness and - injectivity of constructors. - - For example, see \verb|~~/src/HOL/Sum_Type.thy| for the - representation of the primitive sum type as fully-featured datatype. - - \end{description} - - The generated rules for \hyperlink{method.induct}{\mbox{\isa{induct}}} and \hyperlink{method.cases}{\mbox{\isa{cases}}} provide - case names according to the given constructors, while parameters are - named after the types (see also \secref{sec:cases-induct}). - - See \cite{isabelle-HOL} for more details on datatypes, but beware of - the old-style theory syntax being used there! Apart from proper - proof methods for case-analysis and induction, there are also - emulations of ML tactics \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} available, see \secref{sec:hol-induct-tac}; these admit - to refer directly to the internal structure of subgoals (including - internally bound parameters).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We define a type of finite sequences, with slightly different - names than the existing \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a\ list{\isaliteral{22}{\isachardoublequote}}} that is already in \isa{Main}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ {\isaliteral{27}{\isacharprime}}a\ seq\ {\isaliteral{3D}{\isacharequal}}\ Empty\ {\isaliteral{7C}{\isacharbar}}\ Seq\ {\isaliteral{27}{\isacharprime}}a\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ seq{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptext}% -We can now prove some simple lemma by structural induction:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{case}\isamarkupfalse% -\ Empty% -\begin{isamarkuptxt}% -This case can be proved using the simplifier: the freeness - properties of the datatype are already declared as \hyperlink{attribute.simp}{\mbox{\isa{simp}}} rules.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ Empty\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Empty{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}% -\begin{isamarkuptxt}% -The step case is proved similarly.% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ {\isaliteral{28}{\isacharparenleft}}Seq\ y\ ys{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Seq\ y\ ys{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{using}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}Seq\ y\ ys\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ ys{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Here is a more succinct version of the same proof:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Seq\ x\ xs\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ xs{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ xs\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\ simp{\isaliteral{5F}{\isacharunderscore}}all% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{Records \label{sec:hol-record}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In principle, records merely generalize the concept of tuples, where - components may be addressed by labels instead of just position. The - logical infrastructure of records in Isabelle/HOL is slightly more - advanced, though, supporting truly extensible record schemes. This - admits operations that are polymorphic with respect to record - extension, yielding ``object-oriented'' effects like (single) - inheritance. See also \cite{NaraschewskiW-TPHOLs98} for more - details on object-oriented verification and record subtyping in HOL.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Basic concepts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/HOL supports both \emph{fixed} and \emph{schematic} records - at the level of terms and types. The notation is as follows: - - \begin{center} - \begin{tabular}{l|l|l} - & record terms & record types \\ \hline - fixed & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - schematic & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ M{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \end{center} - - \noindent The ASCII representation of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{7C}{\isacharbar}}\ x\ {\isaliteral{3D}{\isacharequal}}\ a\ {\isaliteral{7C}{\isacharbar}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. - - A fixed record \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} has field \isa{x} of value - \isa{a} and field \isa{y} of value \isa{b}. The corresponding - type is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, assuming that \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{22}{\isachardoublequote}}} - and \isa{{\isaliteral{22}{\isachardoublequote}}b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{22}{\isachardoublequote}}}. - - A record scheme like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} contains fields - \isa{x} and \isa{y} as before, but also possibly further fields - as indicated by the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' notation (which is actually part - of the syntax). The improper field ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' of a record - scheme is called the \emph{more part}. Logically it is just a free - variable, which is occasionally referred to as ``row variable'' in - the literature. The more part of a record scheme may be - instantiated by zero or more further components. For example, the - previous scheme may get instantiated to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ m{\isaliteral{27}{\isacharprime}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, where \isa{m{\isaliteral{27}{\isacharprime}}} refers to a different more part. - Fixed records are special instances of record schemes, where - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' is properly terminated by the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ unit{\isaliteral{22}{\isachardoublequote}}} - element. In fact, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} is just an abbreviation - for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}. - - \medskip Two key observations make extensible records in a simply - typed language like HOL work out: - - \begin{enumerate} - - \item the more part is internalized, as a free term or type - variable, - - \item field names are externalized, they cannot be accessed within - the logic as first-class values. - - \end{enumerate} - - \medskip In Isabelle/HOL record types have to be defined explicitly, - fixing their field names and types, and their (optional) parent - record. Afterwards, records may be formed using above syntax, while - obeying the canonical order of fields as given by their declaration. - The record package provides several standard operations like - selectors and updates. The common setup for various generic proof - tools enable succinct reasoning patterns. See also the Isabelle/HOL - tutorial \cite{isabelle-hol-book} for further instructions on using - records in practice.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Record specifications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{record}\hypertarget{command.HOL.record}{\hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{4}{} -\rail@term{\hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}}[] -\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@cr{2} -\rail@bar -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\isa{constdecl}}[] -\rail@nextplus{3} -\rail@endplus -\rail@end -\rail@begin{2}{\isa{constdecl}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.record}{\mbox{\isa{\isacommand{record}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{2B}{\isacharplus}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} defines extensible record type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}}, - derived from the optional parent record \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} by adding new - field components \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} etc. - - The type variables of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} need to be - covered by the (distinct) parameters \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}. Type constructor \isa{t} has to be new, while \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} needs to specify an instance of an existing record type. At - least one new field \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} has to be specified. - Basically, field names need to belong to a unique record. This is - not a real restriction in practice, since fields are qualified by - the record name internally. - - The parent record specification \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} is optional; if omitted - \isa{t} becomes a root record. The hierarchy of all records - declared within a theory context forms a forest structure, i.e.\ a - set of trees starting with a root record each. There is no way to - merge multiple parent records! - - For convenience, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} is made a - type abbreviation for the fixed record type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, likewise is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{5F}{\isacharunderscore}}scheme{\isaliteral{22}{\isachardoublequote}}} made an abbreviation for - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Record operations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Any record definition of the form presented above produces certain - standard operations. Selectors and updates are provided for any - field, including the improper one ``\isa{more}''. There are also - cumulative record constructor functions. To simplify the - presentation below, we assume for now that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} is a root record with fields \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. - - \medskip \textbf{Selectors} and \textbf{updates} are available for - any field (including ``\isa{more}''): - - \begin{matharray}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5F}{\isacharunderscore}}update{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - There is special syntax for application of updates: \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} abbreviates term \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{5F}{\isacharunderscore}}update\ a\ r{\isaliteral{22}{\isachardoublequote}}}. Further notation for - repeated updates is also available: \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} may be written \isa{{\isaliteral{22}{\isachardoublequote}}r{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}. Note that - because of postfix notation the order of fields shown here is - reverse than in the actual term. Since repeated updates are just - function applications, fields may be freely permuted in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ a{\isaliteral{2C}{\isacharcomma}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{2C}{\isacharcomma}}\ z\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}}, as far as logical equality is concerned. - Thus commutativity of independent updates can be proven within the - logic for any two fields, but not as a general theorem. - - \medskip The \textbf{make} operation provides a cumulative record - constructor function: - - \begin{matharray}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \medskip We now reconsider the case of non-root records, which are - derived of some parent. In general, the latter may depend on - another parent as well, resulting in a list of \emph{ancestor - records}. Appending the lists of fields of all ancestors results in - a certain field prefix. The record package automatically takes care - of this by lifting operations over this context of ancestor fields. - Assuming that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} has ancestor - fields \isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}}, - the above record operations will get the following types: - - \medskip - \begin{tabular}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{5F}{\isacharunderscore}}update{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C72686F3E}{\isasymrho}}\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - \noindent Some further operations address the extension aspect of a - derived record scheme specifically: \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} produces a - record fragment consisting of exactly the new fields introduced here - (the result may serve as a more part elsewhere); \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}} - takes a fixed record and adds a given more part; \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} restricts a record scheme to a fixed record. - - \medskip - \begin{tabular}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7A6574613E}{\isasymzeta}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C6C706172723E}{\isasymlparr}}\isaliteral{5C3C5E7665633E}{}\isactrlvec b\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C72686F3E}{\isasymrho}}{\isaliteral{2C}{\isacharcomma}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{5C3C72706172723E}{\isasymrparr}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - \noindent Note that \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}} coincide - for root records.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived rules and proof tools% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The record package proves several results internally, declaring - these facts to appropriate proof tools. This enables users to - reason about record structures quite conveniently. Assume that - \isa{t} is a record type as specified above. - - \begin{enumerate} - - \item Standard conversions for selectors or updates applied to - record constructor terms are made part of the default Simplifier - context; thus proofs by reduction of basic operations merely require - the \hyperlink{method.simp}{\mbox{\isa{simp}}} method without further arguments. These rules - are available as \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}simps{\isaliteral{22}{\isachardoublequote}}}, too. - - \item Selectors applied to updated records are automatically reduced - by an internal simplification procedure, which is also part of the - standard Simplifier setup. - - \item Inject equations of a form analogous to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{27}{\isacharprime}}{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{27}{\isacharprime}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C616E643E}{\isasymand}}\ y\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} are declared to the Simplifier and Classical - Reasoner as \hyperlink{attribute.iff}{\mbox{\isa{iff}}} rules. These rules are available as - \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}iffs{\isaliteral{22}{\isachardoublequote}}}. - - \item The introduction rule for record equality analogous to \isa{{\isaliteral{22}{\isachardoublequote}}x\ r\ {\isaliteral{3D}{\isacharequal}}\ x\ r{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ y\ r\ {\isaliteral{3D}{\isacharequal}}\ y\ r{\isaliteral{27}{\isacharprime}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ r\ {\isaliteral{3D}{\isacharequal}}\ r{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} is declared to the Simplifier, - and as the basic rule context as ``\hyperlink{attribute.intro}{\mbox{\isa{intro}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}''. - The rule is called \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}equality{\isaliteral{22}{\isachardoublequote}}}. - - \item Representations of arbitrary record expressions as canonical - constructor terms are provided both in \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} format (cf.\ the generic proof methods of the same name, - \secref{sec:cases-induct}). Several variations are available, for - fixed records, record schemes, more parts etc. - - The generic proof methods are sufficiently smart to pick the most - sensible rule according to the type of the indicated record - expression: users just need to apply something like ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}cases\ r{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' to a certain proof problem. - - \item The derived record operations \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}make{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}fields{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}extend{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}truncate{\isaliteral{22}{\isachardoublequote}}} are \emph{not} - treated automatically, but usually need to be expanded by hand, - using the collective fact \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{2E}{\isachardot}}defs{\isaliteral{22}{\isachardoublequote}}}. - - \end{enumerate}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -See \verb|~~/src/HOL/ex/Records.thy|, for example.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Adhoc tuples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{attribute}{split\_format}\hypertarget{attribute.HOL.split-format}{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{complete}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{attribute.HOL.split-format}{\mbox{\isa{split{\isaliteral{5F}{\isacharunderscore}}format}}}\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}complete{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} causes - arguments in function applications to be represented canonically - according to their tuple type structure. - - Note that this operation tends to invent funny names for new local - parameters introduced. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Typedef axiomatization \label{sec:hol-typedef}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{typedef}\hypertarget{command.HOL.typedef}{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - A Gordon/HOL-style type definition is a certain axiom scheme that - identifies a new type with a subset of an existing type. More - precisely, the new type is defined by exhibiting an existing type - \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, a set \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ set{\isaliteral{22}{\isachardoublequote}}}, and a theorem that proves - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequote}}}. Thus \isa{A} is a non-empty subset of \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}}, and the new type denotes this subset. New functions are - postulated that establish an isomorphism between the new type and - the subset. In general, the type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} may involve type - variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} which means that the type definition - produces a type constructor \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} depending on - those type arguments. - - The axiomatization can be considered a ``definition'' in the sense - of the particular set-theoretic interpretation of HOL - \cite{pitts93}, where the universe of types is required to be - downwards-closed wrt.\ arbitrary non-empty subsets. Thus genuinely - new types introduced by \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} stay within the range - of HOL models by construction. Note that \indexref{}{command}{type\_synonym}\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}} from Isabelle/Pure merely introduces syntactic - abbreviations, without any logical significance. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}}[] -\rail@endbar -\rail@nont{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}}[] -\rail@end -\rail@begin{3}{\isa{alt{\isaliteral{5F}{\isacharunderscore}}name}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{open}}}[] -\rail@nextbar{2} -\rail@term{\isa{\isakeyword{open}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\isa{abs{\isaliteral{5F}{\isacharunderscore}}type}} -\rail@nont{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{rep{\isaliteral{5F}{\isacharunderscore}}set}} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{morphisms}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ A{\isaliteral{22}{\isachardoublequote}}} - axiomatizes a type definition in the background theory of the - current context, depending on a non-emptiness result of the set - \isa{A} that needs to be proven here. The set \isa{A} may - contain type variables \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} as specified on the LHS, - but no term variables. - - Even though a local theory specification, the newly introduced type - constructor cannot depend on parameters or assumptions of the - context: this is structurally impossible in HOL. In contrast, the - non-emptiness proof may use local assumptions in unusual situations, - which could result in different interpretations in target contexts: - the meaning of the bijection between the representing set \isa{A} - and the new type \isa{t} may then change in different application - contexts. - - By default, \hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}} defines both a type - constructor \isa{t} for the new type, and a term constant \isa{t} for the representing set within the old type. Use the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' option to suppress a separate constant definition - altogether. The injection from type to set is called \isa{Rep{\isaliteral{5F}{\isacharunderscore}}t}, - its inverse \isa{Abs{\isaliteral{5F}{\isacharunderscore}}t}, unless explicit \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} specification provides alternative names. - - The core axiomatization uses the locale predicate \isa{type{\isaliteral{5F}{\isacharunderscore}}definition} as defined in Isabelle/HOL. Various basic - consequences of that are instantiated accordingly, re-using the - locale facts with names derived from the new type constructor. Thus - the generic \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep} is turned into the specific - \isa{{\isaliteral{22}{\isachardoublequote}}Rep{\isaliteral{5F}{\isacharunderscore}}t{\isaliteral{22}{\isachardoublequote}}}, for example. - - Theorems \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep}, \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inverse}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inverse} - provide the most basic characterization as a corresponding - injection/surjection pair (in both directions). The derived rules - \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}inject} and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}inject} provide a more convenient version of - injectivity, suitable for automated proof tools (e.g.\ in - declarations involving \hyperlink{attribute.simp}{\mbox{\isa{simp}}} or \hyperlink{attribute.iff}{\mbox{\isa{iff}}}). - Furthermore, the rules \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}cases}~/ \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Rep{\isaliteral{5F}{\isacharunderscore}}induct}, and \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}cases}~/ - \isa{type{\isaliteral{5F}{\isacharunderscore}}definition{\isaliteral{2E}{\isachardot}}Abs{\isaliteral{5F}{\isacharunderscore}}induct} provide alternative views on - surjectivity. These rules are already declared as set or type rules - for the generic \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} methods, - respectively. - - An alternative name for the set definition (and other derived - entities) may be specified in parentheses; the default is to use - \isa{t} directly. - - \end{description} - - \begin{warn} - If you introduce a new type axiomatically, i.e.\ via \indexref{}{command}{typedecl}\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}} and \indexref{}{command}{axiomatization}\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}, the minimum requirement - is that it has a non-empty model, to avoid immediate collapse of the - HOL logic. Moreover, one needs to demonstrate that the - interpretation of such free-form axiomatizations can coexist with - that of the regular \indexdef{}{command}{typedef}\hypertarget{command.typedef}{\hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}}} scheme, and any extension - that other people might have introduced elsewhere (e.g.\ in HOLCF - \cite{MuellerNvOS99}). - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Type definitions permit the introduction of abstract data - types in a safe way, namely by providing models based on already - existing types. Given some abstract axiomatic description \isa{P} - of a type, this involves two steps: - - \begin{enumerate} - - \item Find an appropriate type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} and subset \isa{A} which - has the desired properties \isa{P}, and make a type definition - based on this representation. - - \item Prove that \isa{P} holds for \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} by lifting \isa{P} - from the representation. - - \end{enumerate} - - You can later forget about the representation and work solely in - terms of the abstract properties \isa{P}. - - \medskip The following trivial example pulls a three-element type - into existence within the formal logical environment of HOL.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{typedef}\isamarkupfalse% -\ three\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}True{\isaliteral{2C}{\isacharcomma}}\ False{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isacommand{definition}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Three\ {\isaliteral{3D}{\isacharequal}}\ Abs{\isaliteral{5F}{\isacharunderscore}}three\ {\isaliteral{28}{\isacharparenleft}}False{\isaliteral{2C}{\isacharcomma}}\ True{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ three{\isaliteral{5F}{\isacharunderscore}}distinct{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}One\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}Two\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp{\isaliteral{5F}{\isacharunderscore}}all\ add{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ three{\isaliteral{5F}{\isacharunderscore}}cases{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{fixes}\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ three\ \isakeyword{obtains}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Two{\isaliteral{22}{\isachardoublequoteclose}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Three{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}cases\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}auto\ simp{\isaliteral{3A}{\isacharcolon}}\ One{\isaliteral{5F}{\isacharunderscore}}def\ Two{\isaliteral{5F}{\isacharunderscore}}def\ Three{\isaliteral{5F}{\isacharunderscore}}def\ Abs{\isaliteral{5F}{\isacharunderscore}}three{\isaliteral{5F}{\isacharunderscore}}inject\ three{\isaliteral{5F}{\isacharunderscore}}def{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -Note that such trivial constructions are better done with - derived specification mechanisms such as \hyperlink{command.datatype}{\mbox{\isa{\isacommand{datatype}}}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ three{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ One{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Two{\isaliteral{27}{\isacharprime}}\ {\isaliteral{7C}{\isacharbar}}\ Three{\isaliteral{27}{\isacharprime}}% -\begin{isamarkuptext}% -This avoids re-doing basic definitions and proofs from the - primitive \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} above.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Functorial structure of types% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{enriched\_type}\hypertarget{command.HOL.enriched-type}{\hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}prefix{\isaliteral{3A}{\isacharcolon}}\ m{\isaliteral{22}{\isachardoublequote}}} allows to - prove and register properties about the functorial structure of type - constructors. These properties then can be used by other packages - to deal with those type constructors in certain type constructions. - Characteristic theorems are noted in the current local theory. By - default, they are prefixed with the base name of the type - constructor, an explicit prefix can be given alternatively. - - The given term \isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{22}{\isachardoublequote}}} is considered as \emph{mapper} for the - corresponding type constructor and must conform to the following - type pattern: - - \begin{matharray}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub k\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \noindent where \isa{t} is the type constructor, \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7665633E}{}\isactrlvec {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} are distinct - type variables free in the local theory and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, - \ldots, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub k{\isaliteral{22}{\isachardoublequote}}} is a subsequence of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \ldots, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Transfer package% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{transfer}\hypertarget{method.HOL.transfer}{\hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{transfer'}\hypertarget{method.HOL.transfer'}{\hyperlink{method.HOL.transfer'}{\mbox{\isa{transfer{\isaliteral{27}{\isacharprime}}}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{transfer\_prover}\hypertarget{method.HOL.transfer-prover}{\hyperlink{method.HOL.transfer-prover}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}prover}}}} & : & \isa{method} \\ - \indexdef{HOL}{attribute}{transfer\_rule}\hypertarget{attribute.HOL.transfer-rule}{\hyperlink{attribute.HOL.transfer-rule}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}rule}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{relator\_eq}\hypertarget{attribute.HOL.relator-eq}{\hyperlink{attribute.HOL.relator-eq}{\mbox{\isa{relator{\isaliteral{5F}{\isacharunderscore}}eq}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{description} - - \item \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} method replaces the current subgoal - with a logically equivalent one that uses different types and - constants. The replacement of types and constants is guided by the - database of transfer rules. Goals are generalized over all free - variables by default; this is necessary for variables whose types - change, but can be overridden for specific variables with e.g. - \isa{{\isaliteral{22}{\isachardoublequote}}transfer\ fixing{\isaliteral{3A}{\isacharcolon}}\ x\ y\ z{\isaliteral{22}{\isachardoublequote}}}. - - \item \hyperlink{method.HOL.transfer'}{\mbox{\isa{transfer{\isaliteral{27}{\isacharprime}}}}} is a variant of \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} that allows replacing a subgoal with one that is - logically stronger (rather than equivalent). For example, a - subgoal involving equality on a quotient type could be replaced - with a subgoal involving equality (instead of the corresponding - equivalence relation) on the underlying raw type. - - \item \hyperlink{method.HOL.transfer-prover}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}prover}}} method assists with proving - a transfer rule for a new constant, provided the constant is - defined in terms of other constants that already have transfer - rules. It should be applied after unfolding the constant - definitions. - - \item \hyperlink{attribute.HOL.transfer-rule}{\mbox{\isa{transfer{\isaliteral{5F}{\isacharunderscore}}rule}}} attribute maintains a - collection of transfer rules, which relate constants at two - different types. Typical transfer rules may relate different type - instances of the same polymorphic constant, or they may relate an - operation on a raw type to a corresponding operation on an - abstract type (quotient or subtype). For example: - - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ B{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ B{\isaliteral{29}{\isacharparenright}}\ map\ map{\isaliteral{22}{\isachardoublequote}}}\\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}cr{\isaliteral{5F}{\isacharunderscore}}int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ cr{\isaliteral{5F}{\isacharunderscore}}int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ cr{\isaliteral{5F}{\isacharunderscore}}int{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2C}{\isacharcomma}}y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}u{\isaliteral{2C}{\isacharcomma}}v{\isaliteral{29}{\isacharparenright}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}x{\isaliteral{2B}{\isacharplus}}u{\isaliteral{2C}{\isacharcomma}}\ y{\isaliteral{2B}{\isacharplus}}v{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ plus{\isaliteral{22}{\isachardoublequote}}} - - Lemmas involving predicates on relations can also be registered - using the same attribute. For example: - - \isa{{\isaliteral{22}{\isachardoublequote}}bi{\isaliteral{5F}{\isacharunderscore}}unique\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ A\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}\ distinct\ distinct{\isaliteral{22}{\isachardoublequote}}}\\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}bi{\isaliteral{5F}{\isacharunderscore}}unique\ A{\isaliteral{3B}{\isacharsemicolon}}\ bi{\isaliteral{5F}{\isacharunderscore}}unique\ B{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ bi{\isaliteral{5F}{\isacharunderscore}}unique\ {\isaliteral{28}{\isacharparenleft}}prod{\isaliteral{5F}{\isacharunderscore}}rel\ A\ B{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - - \item \hyperlink{attribute.HOL.relator-eq}{\mbox{\isa{relator{\isaliteral{5F}{\isacharunderscore}}eq}}} attribute collects identity laws - for relators of various type constructors, e.g. \isa{{\isaliteral{22}{\isachardoublequote}}list{\isaliteral{5F}{\isacharunderscore}}all{\isadigit{2}}\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}op\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. The \hyperlink{method.HOL.transfer}{\mbox{\isa{transfer}}} method uses these - lemmas to infer transfer rules for non-polymorphic constants on - the fly. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Lifting package% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{setup\_lifting}\hypertarget{command.HOL.setup-lifting}{\hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{lift\_definition}\hypertarget{command.HOL.lift-definition}{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{print\_quotmaps}\hypertarget{command.HOL.print-quotmaps}{\hyperlink{command.HOL.print-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{print\_quotients}\hypertarget{command.HOL.print-quotients}{\hyperlink{command.HOL.print-quotients}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotients}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{attribute}{quot\_map}\hypertarget{attribute.HOL.quot-map}{\hyperlink{attribute.HOL.quot-map}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}map}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{invariant\_commute}\hypertarget{attribute.HOL.invariant-commute}{\hyperlink{attribute.HOL.invariant-commute}{\mbox{\isa{invariant{\isaliteral{5F}{\isacharunderscore}}commute}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{5}{} -\rail@term{\hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}abs{\isaliteral{5F}{\isacharunderscore}}code}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{railoutput} -\rail@begin{4}{} -\rail@term{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@cr{3} -\rail@term{\isa{is}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} Sets up the Lifting package - to work with a user-defined type. The user must provide either a - quotient theorem \isa{{\isaliteral{22}{\isachardoublequote}}Quotient\ R\ Abs\ Rep\ T{\isaliteral{22}{\isachardoublequote}}} or a - type_definition theorem \isa{{\isaliteral{22}{\isachardoublequote}}type{\isaliteral{5F}{\isacharunderscore}}definition\ Rep\ Abs\ A{\isaliteral{22}{\isachardoublequote}}}. The - package configures transfer rules for equality and quantifiers on - the type, and sets up the \indexdef{HOL}{command}{lift\_definition}\hypertarget{command.HOL.lift-definition}{\hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}}} - command to work with the type. In the case of a quotient theorem, - an optional theorem \isa{{\isaliteral{22}{\isachardoublequote}}reflp\ R{\isaliteral{22}{\isachardoublequote}}} can be provided as a second - argument. This allows the package to generate stronger transfer - rules. - - \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} is called automatically if a - quotient type is defined by the command \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} from the Quotient package. - - If \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} is called with a - type_definition theorem, the abstract type implicitly defined by - the theorem is declared as an abstract type in the code - generator. This allows \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} to - register (generated) code certificate theorems as abstract code - equations in the code generator. The option \isa{{\isaliteral{22}{\isachardoublequote}}no{\isaliteral{5F}{\isacharunderscore}}abs{\isaliteral{5F}{\isacharunderscore}}code{\isaliteral{22}{\isachardoublequote}}} - of the command \hyperlink{command.HOL.setup-lifting}{\mbox{\isa{\isacommand{setup{\isaliteral{5F}{\isacharunderscore}}lifting}}}} can turn off that - behavior and causes that code certificate theorems generated by - \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} are not registred as abstract - code equations. - - \item \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ is\ t{\isaliteral{22}{\isachardoublequote}}} - Defines a new function \isa{f} with an abstract type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} - in terms of a corresponding operation \isa{t} on a - representation type. The term \isa{t} doesn't have to be - necessarily a constant but it can be any term. - - Users must discharge a respectfulness proof obligation when each - constant is defined. For a type copy, i.e. a typedef with \isa{UNIV}, the proof is discharged automatically. The obligation is - presented in a user-friendly, readable form. A respectfulness - theorem in the standard format \isa{f{\isaliteral{2E}{\isachardot}}rsp} and a transfer rule - \isa{f{\isaliteral{2E}{\isachardot}}tranfer} for the Transfer package are generated by the - package. - - Integration with code_abstype: For typedefs (e.g. subtypes - corresponding to a datatype invariant, such as dlist), \hyperlink{command.HOL.lift-definition}{\mbox{\isa{\isacommand{lift{\isaliteral{5F}{\isacharunderscore}}definition}}}} generates a code certificate theorem - \isa{f{\isaliteral{2E}{\isachardot}}rep{\isaliteral{5F}{\isacharunderscore}}eq} and sets up code generation for each constant. - - \item \hyperlink{command.HOL.print-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}} prints stored quotient map - theorems. - - \item \hyperlink{command.HOL.print-quotients}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotients}}}} prints stored quotient - theorems. - - \item \hyperlink{attribute.HOL.quot-map}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}map}}} registers a quotient map - theorem. For examples see \verb|~~/src/HOL/Library/Quotient_List.thy| or other Quotient_*.thy - files. - - \item \hyperlink{attribute.HOL.invariant-commute}{\mbox{\isa{invariant{\isaliteral{5F}{\isacharunderscore}}commute}}} registers a theorem which - shows a relationship between the constant \isa{Lifting{\isaliteral{2E}{\isachardot}}invariant} (used for internal encoding of proper subtypes) - and a relator. Such theorems allows the package to hide \isa{Lifting{\isaliteral{2E}{\isachardot}}invariant} from a user in a user-readable form of a - respectfulness theorem. For examples see \verb|~~/src/HOL/Library/Quotient_List.thy| or other Quotient_*.thy - files. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Quotient types% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{quotient\_type}\hypertarget{command.HOL.quotient-type}{\hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{quotient\_definition}\hypertarget{command.HOL.quotient-definition}{\hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{print\_quotmapsQ3}\hypertarget{command.HOL.print-quotmapsQ3}{\hyperlink{command.HOL.print-quotmapsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmapsQ{\isadigit{3}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{print\_quotientsQ3}\hypertarget{command.HOL.print-quotientsQ3}{\hyperlink{command.HOL.print-quotientsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotientsQ{\isadigit{3}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{command}{print\_quotconsts}\hypertarget{command.HOL.print-quotconsts}{\hyperlink{command.HOL.print-quotconsts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotconsts}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}}\\ - \indexdef{HOL}{method}{lifting}\hypertarget{method.HOL.lifting}{\hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{lifting\_setup}\hypertarget{method.HOL.lifting-setup}{\hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{descending}\hypertarget{method.HOL.descending}{\hyperlink{method.HOL.descending}{\mbox{\isa{descending}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{descending\_setup}\hypertarget{method.HOL.descending-setup}{\hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{partiality\_descending}\hypertarget{method.HOL.partiality-descending}{\hyperlink{method.HOL.partiality-descending}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{partiality\_descending\_setup}\hypertarget{method.HOL.partiality-descending-setup}{\hyperlink{method.HOL.partiality-descending-setup}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending{\isaliteral{5F}{\isacharunderscore}}setup}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{regularize}\hypertarget{method.HOL.regularize}{\hyperlink{method.HOL.regularize}{\mbox{\isa{regularize}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{injection}\hypertarget{method.HOL.injection}{\hyperlink{method.HOL.injection}{\mbox{\isa{injection}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{cleaning}\hypertarget{method.HOL.cleaning}{\hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}}} & : & \isa{method} \\ - \indexdef{HOL}{attribute}{quot\_thm}\hypertarget{attribute.HOL.quot-thm}{\hyperlink{attribute.HOL.quot-thm}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}thm}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{quot\_lifted}\hypertarget{attribute.HOL.quot-lifted}{\hyperlink{attribute.HOL.quot-lifted}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}lifted}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{quot\_respect}\hypertarget{attribute.HOL.quot-respect}{\hyperlink{attribute.HOL.quot-respect}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}respect}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{quot\_preserve}\hypertarget{attribute.HOL.quot-preserve}{\hyperlink{attribute.HOL.quot-preserve}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}preserve}}}} & : & \isa{attribute} \\ - \end{matharray} - - The quotient package defines a new quotient type given a raw type - and a partial equivalence relation. It also includes automation for - transporting definitions and theorems. It can automatically produce - definitions and theorems on the quotient type, given the - corresponding constants and facts on the raw type. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}}}[] -\rail@plus -\rail@nont{\isa{spec}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{8}{\isa{spec}} -\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@cr{3} -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@term{\isa{{\isaliteral{2F}{\isacharslash}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{partial}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@cr{6} -\rail@bar -\rail@nextbar{7} -\rail@term{\isa{\isakeyword{morphisms}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{railoutput} -\rail@begin{4}{} -\rail@term{\hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{constdecl}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@term{\isa{is}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{2}{\isa{constdecl}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} defines quotient types. The - injection from a quotient type to a raw type is called \isa{rep{\isaliteral{5F}{\isacharunderscore}}t}, its inverse \isa{abs{\isaliteral{5F}{\isacharunderscore}}t} unless explicit \hyperlink{keyword.HOL.morphisms}{\mbox{\isa{\isakeyword{morphisms}}}} specification provides alternative names. \hyperlink{command.HOL.quotient-type}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}type}}}} requires the user to prove that the relation - is an equivalence relation (predicate \isa{equivp}), unless the - user specifies explicitely \isa{partial} in which case the - obligation is \isa{part{\isaliteral{5F}{\isacharunderscore}}equivp}. A quotient defined with \isa{partial} is weaker in the sense that less things can be proved - automatically. - - \item \hyperlink{command.HOL.quotient-definition}{\mbox{\isa{\isacommand{quotient{\isaliteral{5F}{\isacharunderscore}}definition}}}} defines a constant on - the quotient type. - - \item \hyperlink{command.HOL.print-quotmapsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmapsQ{\isadigit{3}}}}}} prints quotient map - functions. - - \item \hyperlink{command.HOL.print-quotientsQ3}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotientsQ{\isadigit{3}}}}}} prints quotients. - - \item \hyperlink{command.HOL.print-quotconsts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotconsts}}}} prints quotient constants. - - \item \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} and \hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}} - methods match the current goal with the given raw theorem to be - lifted producing three new subgoals: regularization, injection and - cleaning subgoals. \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} tries to apply the - heuristics for automatically solving these three subgoals and - leaves only the subgoals unsolved by the heuristics to the user as - opposed to \hyperlink{method.HOL.lifting-setup}{\mbox{\isa{lifting{\isaliteral{5F}{\isacharunderscore}}setup}}} which leaves the three - subgoals unsolved. - - \item \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} and \hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}} try to guess a raw statement that would lift - to the current subgoal. Such statement is assumed as a new subgoal - and \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} continues in the same way as - \hyperlink{method.HOL.lifting}{\mbox{\isa{lifting}}} does. \hyperlink{method.HOL.descending}{\mbox{\isa{descending}}} tries - to solve the arising regularization, injection and cleaning - subgoals with the analoguous method \hyperlink{method.HOL.descending-setup}{\mbox{\isa{descending{\isaliteral{5F}{\isacharunderscore}}setup}}} which leaves the four unsolved subgoals. - - \item \hyperlink{method.HOL.partiality-descending}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending}}} finds the regularized - theorem that would lift to the current subgoal, lifts it and - leaves as a subgoal. This method can be used with partial - equivalence quotients where the non regularized statements would - not be true. \hyperlink{method.HOL.partiality-descending-setup}{\mbox{\isa{partiality{\isaliteral{5F}{\isacharunderscore}}descending{\isaliteral{5F}{\isacharunderscore}}setup}}} leaves - the injection and cleaning subgoals unchanged. - - \item \hyperlink{method.HOL.regularize}{\mbox{\isa{regularize}}} applies the regularization - heuristics to the current subgoal. - - \item \hyperlink{method.HOL.injection}{\mbox{\isa{injection}}} applies the injection heuristics - to the current goal using the stored quotient respectfulness - theorems. - - \item \hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}} applies the injection cleaning - heuristics to the current subgoal using the stored quotient - preservation theorems. - - \item \hyperlink{attribute.HOL.quot-lifted}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}lifted}}} attribute tries to - automatically transport the theorem to the quotient type. - The attribute uses all the defined quotients types and quotient - constants often producing undesired results or theorems that - cannot be lifted. - - \item \hyperlink{attribute.HOL.quot-respect}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}respect}}} and \hyperlink{attribute.HOL.quot-preserve}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}preserve}}} attributes declare a theorem as a respectfulness - and preservation theorem respectively. These are stored in the - local theory store and used by the \hyperlink{method.HOL.injection}{\mbox{\isa{injection}}} - and \hyperlink{method.HOL.cleaning}{\mbox{\isa{cleaning}}} methods respectively. - - \item \hyperlink{attribute.HOL.quot-thm}{\mbox{\isa{quot{\isaliteral{5F}{\isacharunderscore}}thm}}} declares that a certain theorem - is a quotient extension theorem. Quotient extension theorems - allow for quotienting inside container types. Given a polymorphic - type that serves as a container, a map function defined for this - container using \hyperlink{command.HOL.enriched-type}{\mbox{\isa{\isacommand{enriched{\isaliteral{5F}{\isacharunderscore}}type}}}} and a relation - map defined for for the container type, the quotient extension - theorem should be \isa{{\isaliteral{22}{\isachardoublequote}}Quotient{\isadigit{3}}\ R\ Abs\ Rep\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Quotient{\isadigit{3}}\ {\isaliteral{28}{\isacharparenleft}}rel{\isaliteral{5F}{\isacharunderscore}}map\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}map\ Abs{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}map\ Rep{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. Quotient extension theorems - are stored in a database and are used all the steps of lifting - theorems. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Coercive subtyping% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{attribute}{coercion}\hypertarget{attribute.HOL.coercion}{\hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{coercion\_enabled}\hypertarget{attribute.HOL.coercion-enabled}{\hyperlink{attribute.HOL.coercion-enabled}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}enabled}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{coercion\_map}\hypertarget{attribute.HOL.coercion-map}{\hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}} & : & \isa{attribute} \\ - \end{matharray} - - Coercive subtyping allows the user to omit explicit type - conversions, also called \emph{coercions}. Type inference will add - them as necessary when parsing a term. See - \cite{traytel-berghofer-nipkow-2011} for details. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{attribute.HOL.coercion}{\mbox{\isa{coercion}}}~\isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} registers a new - coercion function \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} where \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} are type constructors without arguments. Coercions are - composed by the inference algorithm if needed. Note that the type - inference algorithm is complete only if the registered coercions - form a lattice. - - \item \hyperlink{attribute.HOL.coercion-map}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}map}}}~\isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} registers a - new map function to lift coercions through type constructors. The - function \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} must conform to the following type pattern - - \begin{matharray}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}map{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ f\isaliteral{5C3C5E697375623E}{}\isactrlisub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - where \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} is a type constructor and \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} is of type - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}}. Registering a map function - overwrites any existing map function for this particular type - constructor. - - \item \hyperlink{attribute.HOL.coercion-enabled}{\mbox{\isa{coercion{\isaliteral{5F}{\isacharunderscore}}enabled}}} enables the coercion - inference algorithm. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Arithmetic proof support% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{arith}\hypertarget{method.HOL.arith}{\hyperlink{method.HOL.arith}{\mbox{\isa{arith}}}} & : & \isa{method} \\ - \indexdef{HOL}{attribute}{arith}\hypertarget{attribute.HOL.arith}{\hyperlink{attribute.HOL.arith}{\mbox{\isa{arith}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{arith\_split}\hypertarget{attribute.HOL.arith-split}{\hyperlink{attribute.HOL.arith-split}{\mbox{\isa{arith{\isaliteral{5F}{\isacharunderscore}}split}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{description} - - \item \hyperlink{method.HOL.arith}{\mbox{\isa{arith}}} decides linear arithmetic problems (on - types \isa{nat}, \isa{int}, \isa{real}). Any current facts - are inserted into the goal before running the procedure. - - \item \hyperlink{attribute.HOL.arith}{\mbox{\isa{arith}}} declares facts that are supplied to - the arithmetic provers implicitly. - - \item \hyperlink{attribute.HOL.arith-split}{\mbox{\isa{arith{\isaliteral{5F}{\isacharunderscore}}split}}} attribute declares case split - rules to be expanded before \hyperlink{method.HOL.arith}{\mbox{\isa{arith}}} is invoked. - - \end{description} - - Note that a simpler (but faster) arithmetic prover is already - invoked by the Simplifier.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Intuitionistic proof search% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{iprover}\hypertarget{method.HOL.iprover}{\hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.rulemod}{\mbox{\isa{rulemod}}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}} performs intuitionistic proof search, - depending on specifically declared rules from the context, or given - as explicit arguments. Chained facts are inserted into the goal - before commencing proof search. - - Rules need to be classified as \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}, - \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, or \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}}; here the - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' indicator refers to ``safe'' rules, which may be - applied aggressively (without considering back-tracking later). - Rules declared with ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' are ignored in proof search (the - single-step \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method still observes these). An - explicit weight annotation may be given as well; otherwise the - number of rule premises will be taken into account here. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Model Elimination and Resolution% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{meson}\hypertarget{method.HOL.meson}{\hyperlink{method.HOL.meson}{\mbox{\isa{meson}}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{metis}\hypertarget{method.HOL.metis}{\hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.meson}{\mbox{\isa{meson}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@term{\isa{partial{\isaliteral{5F}{\isacharunderscore}}types}}[] -\rail@nextbar{2} -\rail@term{\isa{full{\isaliteral{5F}{\isacharunderscore}}types}}[] -\rail@nextbar{3} -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}types}}[] -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.meson}{\mbox{\isa{meson}}} implements Loveland's model elimination - procedure \cite{loveland-78}. See \verb|~~/src/HOL/ex/Meson_Test.thy| for examples. - - \item \hyperlink{method.HOL.metis}{\mbox{\isa{metis}}} combines ordered resolution and ordered - paramodulation to find first-order (or mildly higher-order) proofs. - The first optional argument specifies a type encoding; see the - Sledgehammer manual \cite{isabelle-sledgehammer} for details. The - directory \verb|~~/src/HOL/Metis_Examples| contains several small - theories developed to a large extent using \hyperlink{method.HOL.metis}{\mbox{\isa{metis}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Coherent Logic% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{method}{coherent}\hypertarget{method.HOL.coherent}{\hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}}} & : & \isa{method} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}} solves problems of \emph{Coherent - Logic} \cite{Bezem-Coquand:2005}, which covers applications in - confluence theory, lattice theory and projective geometry. See - \verb|~~/src/HOL/ex/Coherent.thy| for some examples. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proving propositions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In addition to the standard proof methods, a number of diagnosis - tools search for proofs and provide an Isar proof snippet on success. - These tools are available via the following commands. - - \begin{matharray}{rcl} - \indexdef{HOL}{command}{solve\_direct}\hypertarget{command.HOL.solve-direct}{\hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{try}\hypertarget{command.HOL.try}{\hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{try0}\hypertarget{command.HOL.try0}{\hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{sledgehammer}\hypertarget{command.HOL.sledgehammer}{\hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{sledgehammer\_params}\hypertarget{command.HOL.sledgehammer-params}{\hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}}}[] -\rail@end -\rail@begin{6}{} -\rail@term{\hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@bar -\rail@term{\isa{simp}}[] -\rail@nextbar{2} -\rail@term{\isa{intro}}[] -\rail@nextbar{3} -\rail@term{\isa{elim}}[] -\rail@nextbar{4} -\rail@term{\isa{dest}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{5} -\rail@endplus -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\isa{args}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{facts}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\isa{args}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{args}} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\isa{value}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@end -\rail@begin{5}{\isa{facts}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@bar -\rail@nextbar{2} -\rail@bar -\rail@term{\isa{add}}[] -\rail@nextbar{3} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - % FIXME check args "value" - - \begin{description} - - \item \hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}} checks whether the current - subgoals can be solved directly by an existing theorem. Duplicate - lemmas can be detected in this way. - - \item \hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}} attempts to prove a subgoal - using a combination of standard proof methods (\hyperlink{method.auto}{\mbox{\isa{auto}}}, - \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.blast}{\mbox{\isa{blast}}}, etc.). Additional facts supplied - via \isa{{\isaliteral{22}{\isachardoublequote}}simp{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}intro{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}elim{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}, and \isa{{\isaliteral{22}{\isachardoublequote}}dest{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}} are passed to the appropriate proof methods. - - \item \hyperlink{command.HOL.try}{\mbox{\isa{\isacommand{try}}}} attempts to prove or disprove a subgoal - using a combination of provers and disprovers (\hyperlink{command.HOL.solve-direct}{\mbox{\isa{\isacommand{solve{\isaliteral{5F}{\isacharunderscore}}direct}}}}, \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}, \hyperlink{command.HOL.try0}{\mbox{\isa{\isacommand{try{\isadigit{0}}}}}}, \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}}, \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}). - - \item \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}} attempts to prove a subgoal - using external automatic provers (resolution provers and SMT - solvers). See the Sledgehammer manual \cite{isabelle-sledgehammer} - for details. - - \item \hyperlink{command.HOL.sledgehammer-params}{\mbox{\isa{\isacommand{sledgehammer{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.sledgehammer}{\mbox{\isa{\isacommand{sledgehammer}}}} configuration options persistently. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Checking and refuting propositions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Identifying incorrect propositions usually involves evaluation of - particular assignments and systematic counterexample search. This - is supported by the following commands. - - \begin{matharray}{rcl} - \indexdef{HOL}{command}{value}\hypertarget{command.HOL.value}{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{values}\hypertarget{command.HOL.values}{\hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{quickcheck}\hypertarget{command.HOL.quickcheck}{\hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{refute}\hypertarget{command.HOL.refute}{\hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{nitpick}\hypertarget{command.HOL.nitpick}{\hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{quickcheck\_params}\hypertarget{command.HOL.quickcheck-params}{\hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{refute\_params}\hypertarget{command.HOL.refute-params}{\hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{nitpick\_params}\hypertarget{command.HOL.nitpick-params}{\hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{quickcheck\_generator}\hypertarget{command.HOL.quickcheck-generator}{\hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{find\_unused\_assms}\hypertarget{command.HOL.find-unused-assms}{\hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{modes}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{modes}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\isa{args}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\isa{args}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@end -\rail@begin{4}{} -\rail@term{\hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@cr{2} -\rail@term{\isa{operations{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{3} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{modes}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\isa{args}} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\isa{value}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@end -\end{railoutput} - % FIXME check "value" - - \begin{description} - - \item \hyperlink{command.HOL.value}{\mbox{\isa{\isacommand{value}}}}~\isa{t} evaluates and prints a - term; optionally \isa{modes} can be specified, which are appended - to the current print mode; see \secref{sec:print-modes}. - Internally, the evaluation is performed by registered evaluators, - which are invoked sequentially until a result is returned. - Alternatively a specific evaluator can be selected using square - brackets; typical evaluators use the current set of code equations - to normalize and include \isa{simp} for fully symbolic evaluation - using the simplifier, \isa{nbe} for \emph{normalization by - evaluation} and \emph{code} for code generation in SML. - - \item \hyperlink{command.HOL.values}{\mbox{\isa{\isacommand{values}}}}~\isa{t} enumerates a set - comprehension by evaluation and prints its values up to the given - number of solutions; optionally \isa{modes} can be specified, - which are appended to the current print mode; see - \secref{sec:print-modes}. - - \item \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} tests the current goal for - counterexamples using a series of assignments for its free - variables; by default the first subgoal is tested, an other can be - selected explicitly using an optional goal index. Assignments can - be chosen exhausting the search space upto a given size, or using a - fixed number of random assignments in the search space, or exploring - the search space symbolically using narrowing. By default, - quickcheck uses exhaustive testing. A number of configuration - options are supported for \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}}, notably: - - \begin{description} - - \item[\isa{tester}] specifies which testing approach to apply. - There are three testers, \isa{exhaustive}, \isa{random}, and - \isa{narrowing}. An unknown configuration option is treated as - an argument to tester, making \isa{{\isaliteral{22}{\isachardoublequote}}tester\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} optional. When - multiple testers are given, these are applied in parallel. If no - tester is specified, quickcheck uses the testers that are set - active, i.e., configurations \hyperlink{attribute.quickcheck-exhaustive-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}exhaustive{\isaliteral{5F}{\isacharunderscore}}active}}}, \hyperlink{attribute.quickcheck-random-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}random{\isaliteral{5F}{\isacharunderscore}}active}}}, \hyperlink{attribute.quickcheck-narrowing-active}{\mbox{\isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}narrowing{\isaliteral{5F}{\isacharunderscore}}active}}} are set to true. - - \item[\isa{size}] specifies the maximum size of the search space - for assignment values. - - \item[\isa{genuine{\isaliteral{5F}{\isacharunderscore}}only}] sets quickcheck only to return genuine - counterexample, but not potentially spurious counterexamples due - to underspecified functions. - - \item[\isa{abort{\isaliteral{5F}{\isacharunderscore}}potential}] sets quickcheck to abort once it - found a potentially spurious counterexample and to not continue - to search for a further genuine counterexample. - For this option to be effective, the \isa{genuine{\isaliteral{5F}{\isacharunderscore}}only} option - must be set to false. - - \item[\isa{eval}] takes a term or a list of terms and evaluates - these terms under the variable assignment found by quickcheck. - This option is currently only supported by the default - (exhaustive) tester. - - \item[\isa{iterations}] sets how many sets of assignments are - generated for each particular size. - - \item[\isa{no{\isaliteral{5F}{\isacharunderscore}}assms}] specifies whether assumptions in - structured proofs should be ignored. - - \item[\isa{locale}] specifies how to process conjectures in - a locale context, i.e., they can be interpreted or expanded. - The option is a whitespace-separated list of the two words - \isa{interpret} and \isa{expand}. The list determines the - order they are employed. The default setting is to first use - interpretations and then test the expanded conjecture. - The option is only provided as attribute declaration, but not - as parameter to the command. - - \item[\isa{timeout}] sets the time limit in seconds. - - \item[\isa{default{\isaliteral{5F}{\isacharunderscore}}type}] sets the type(s) generally used to - instantiate type variables. - - \item[\isa{report}] if set quickcheck reports how many tests - fulfilled the preconditions. - - \item[\isa{use{\isaliteral{5F}{\isacharunderscore}}subtype}] if set quickcheck automatically lifts - conjectures to registered subtypes if possible, and tests the - lifted conjecture. - - \item[\isa{quiet}] if set quickcheck does not output anything - while testing. - - \item[\isa{verbose}] if set quickcheck informs about the current - size and cardinality while testing. - - \item[\isa{expect}] can be used to check if the user's - expectation was met (\isa{no{\isaliteral{5F}{\isacharunderscore}}expectation}, \isa{no{\isaliteral{5F}{\isacharunderscore}}counterexample}, or \isa{counterexample}). - - \end{description} - - These option can be given within square brackets. - - \item \hyperlink{command.HOL.quickcheck-params}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.quickcheck}{\mbox{\isa{\isacommand{quickcheck}}}} configuration options persistently. - - \item \hyperlink{command.HOL.quickcheck-generator}{\mbox{\isa{\isacommand{quickcheck{\isaliteral{5F}{\isacharunderscore}}generator}}}} creates random and - exhaustive value generators for a given type and operations. It - generates values by using the operations as if they were - constructors of that type. - - \item \hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}} tests the current goal for - counterexamples using a reduction to SAT. The following - configuration options are supported: - - \begin{description} - - \item[\isa{minsize}] specifies the minimum size (cardinality) of - the models to search for. - - \item[\isa{maxsize}] specifies the maximum size (cardinality) of - the models to search for. Nonpositive values mean \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E66696E6974793E}{\isasyminfinity}}{\isaliteral{22}{\isachardoublequote}}}. - - \item[\isa{maxvars}] specifies the maximum number of Boolean - variables to use when transforming the term into a propositional - formula. Nonpositive values mean \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E66696E6974793E}{\isasyminfinity}}{\isaliteral{22}{\isachardoublequote}}}. - - \item[\isa{satsolver}] specifies the SAT solver to use. - - \item[\isa{no{\isaliteral{5F}{\isacharunderscore}}assms}] specifies whether assumptions in - structured proofs should be ignored. - - \item[\isa{maxtime}] sets the time limit in seconds. - - \item[\isa{expect}] can be used to check if the user's - expectation was met (\isa{genuine}, \isa{potential}, \isa{none}, or \isa{unknown}). - - \end{description} - - These option can be given within square brackets. - - \item \hyperlink{command.HOL.refute-params}{\mbox{\isa{\isacommand{refute{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.refute}{\mbox{\isa{\isacommand{refute}}}} configuration options persistently. - - \item \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}} tests the current goal for - counterexamples using a reduction to first-order relational - logic. See the Nitpick manual \cite{isabelle-nitpick} for details. - - \item \hyperlink{command.HOL.nitpick-params}{\mbox{\isa{\isacommand{nitpick{\isaliteral{5F}{\isacharunderscore}}params}}}} changes \hyperlink{command.HOL.nitpick}{\mbox{\isa{\isacommand{nitpick}}}} configuration options persistently. - - \item \hyperlink{command.HOL.find-unused-assms}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}unused{\isaliteral{5F}{\isacharunderscore}}assms}}}} finds potentially superfluous - assumptions in theorems using quickcheck. - It takes the theory name to be checked for superfluous assumptions as - optional argument. If not provided, it checks the current theory. - Options to the internal quickcheck invocations can be changed with - common configuration declarations. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Unstructured case analysis and induction \label{sec:hol-induct-tac}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following tools of Isabelle/HOL support cases analysis and - induction in unstructured tactic scripts; see also - \secref{sec:cases-induct} for proper Isar versions of similar ideas. - - \begin{matharray}{rcl} - \indexdef{HOL}{method}{case\_tac}\hypertarget{method.HOL.case-tac}{\hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{induct\_tac}\hypertarget{method.HOL.induct-tac}{\hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{HOL}{method}{ind\_cases}\hypertarget{method.HOL.ind-cases}{\hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\ - \indexdef{HOL}{command}{inductive\_cases}\hypertarget{command.HOL.inductive-cases}{\hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{2} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}}[] -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{1}{\isa{rule}} -\rail@term{\isa{rule}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.HOL.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}} and \hyperlink{method.HOL.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} admit - to reason about inductive types. Rules are selected according to - the declarations by the \hyperlink{attribute.cases}{\mbox{\isa{cases}}} and \hyperlink{attribute.induct}{\mbox{\isa{induct}}} - attributes, cf.\ \secref{sec:cases-induct}. The \hyperlink{command.HOL.datatype}{\mbox{\isa{\isacommand{datatype}}}} package already takes care of this. - - These unstructured tactics feature both goal addressing and dynamic - instantiation. Note that named rule cases are \emph{not} provided - as would be by the proper \hyperlink{method.cases}{\mbox{\isa{cases}}} and \hyperlink{method.induct}{\mbox{\isa{induct}}} proof - methods (see \secref{sec:cases-induct}). Unlike the \hyperlink{method.induct}{\mbox{\isa{induct}}} method, \hyperlink{method.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}} does not handle structured rule - statements, only the compact object-logic conclusion of the subgoal - being addressed. - - \item \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} and \hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}} provide an interface to the internal \verb|mk_cases| operation. Rules are simplified in an unrestricted - forward manner. - - While \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} is a proof method to apply the - result immediately as elimination rules, \hyperlink{command.HOL.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}} provides case split theorems at the theory level - for later use. The \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} argument of the \hyperlink{method.HOL.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}} method allows to specify a list of variables that should - be generalized before applying the resulting rule. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Executable code% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For validation purposes, it is often useful to \emph{execute} - specifications. In principle, execution could be simulated by - Isabelle's inference kernel, i.e. by a combination of resolution and - simplification. Unfortunately, this approach is rather inefficient. - A more efficient way of executing specifications is to translate - them into a functional programming language such as ML. - - Isabelle provides a generic framework to support code generation - from executable specifications. Isabelle/HOL instantiates these - mechanisms in a way that is amenable to end-user applications. Code - can be generated for functional programs (including overloading - using type classes) targeting SML \cite{SML}, OCaml \cite{OCaml}, - Haskell \cite{haskell-revised-report} and Scala - \cite{scala-overview-tech-report}. Conceptually, code generation is - split up in three steps: \emph{selection} of code theorems, - \emph{translation} into an abstract executable view and - \emph{serialization} to a specific \emph{target language}. - Inductive specifications can be executed using the predicate - compiler which operates within HOL. See \cite{isabelle-codegen} for - an introduction. - - \begin{matharray}{rcl} - \indexdef{HOL}{command}{export\_code}\hypertarget{command.HOL.export-code}{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{code}\hypertarget{attribute.HOL.code}{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{command}{code\_abort}\hypertarget{command.HOL.code-abort}{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_datatype}\hypertarget{command.HOL.code-datatype}{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{print\_codesetup}\hypertarget{command.HOL.print-codesetup}{\hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{attribute}{code\_unfold}\hypertarget{attribute.HOL.code-unfold}{\hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{code\_post}\hypertarget{attribute.HOL.code-post}{\hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{attribute}{code\_abbrev}\hypertarget{attribute.HOL.code-abbrev}{\hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}}} & : & \isa{attribute} \\ - \indexdef{HOL}{command}{print\_codeproc}\hypertarget{command.HOL.print-codeproc}{\hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codeproc}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_thms}\hypertarget{command.HOL.code-thms}{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_deps}\hypertarget{command.HOL.code-deps}{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_const}\hypertarget{command.HOL.code-const}{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_type}\hypertarget{command.HOL.code-type}{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_class}\hypertarget{command.HOL.code-class}{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_instance}\hypertarget{command.HOL.code-instance}{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_reserved}\hypertarget{command.HOL.code-reserved}{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_monad}\hypertarget{command.HOL.code-monad}{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_include}\hypertarget{command.HOL.code-include}{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_modulename}\hypertarget{command.HOL.code-modulename}{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_reflect}\hypertarget{command.HOL.code-reflect}{\hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{code\_pred}\hypertarget{command.HOL.code-pred}{\hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - \end{matharray} - - \begin{railoutput} -\rail@begin{11}{} -\rail@term{\hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}}}[] -\rail@plus -\rail@nont{\isa{constexpr}}[] -\rail@nextplus{1} -\rail@endplus -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@plus -\rail@term{\isa{\isakeyword{in}}}[] -\rail@nont{\isa{target}}[] -\rail@bar -\rail@nextbar{5} -\rail@term{\isa{\isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@endbar -\rail@cr{7} -\rail@bar -\rail@nextbar{8} -\rail@term{\isa{\isakeyword{file}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextbar{9} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@endbar -\rail@endbar -\rail@bar -\rail@nextbar{8} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{args}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nextplus{10} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{1}{\isa{const}} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{3}{\isa{constexpr}} -\rail@bar -\rail@nont{\isa{const}}[] -\rail@nextbar{1} -\rail@term{\isa{name{\isaliteral{2E}{\isachardot}}{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\isa{typeconstructor}} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\rail@begin{1}{\isa{class}} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\rail@begin{4}{\isa{target}} -\rail@bar -\rail@term{\isa{SML}}[] -\rail@nextbar{1} -\rail@term{\isa{OCaml}}[] -\rail@nextbar{2} -\rail@term{\isa{Haskell}}[] -\rail@nextbar{3} -\rail@term{\isa{Scala}}[] -\rail@endbar -\rail@end -\rail@begin{4}{} -\rail@term{\hyperlink{attribute.HOL.code}{\mbox{\isa{code}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@bar -\rail@term{\isa{del}}[] -\rail@nextbar{2} -\rail@term{\isa{abstype}}[] -\rail@nextbar{3} -\rail@term{\isa{abstract}}[] -\rail@endbar -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}}}[] -\rail@plus -\rail@nont{\isa{const}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}}}[] -\rail@plus -\rail@nont{\isa{const}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}}}[] -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\isa{constexpr}}[] -\rail@nextplus{2} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\isa{constexpr}}[] -\rail@nextplus{2} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{7}{} -\rail@term{\hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}}}[] -\rail@plus -\rail@nont{\isa{const}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@cr{3} -\rail@plus -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{target}}[] -\rail@plus -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{syntax}}[] -\rail@endbar -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextplus{6} -\rail@endplus -\rail@end -\rail@begin{7}{} -\rail@term{\hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}}}[] -\rail@plus -\rail@nont{\isa{typeconstructor}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@cr{3} -\rail@plus -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{target}}[] -\rail@plus -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{syntax}}[] -\rail@endbar -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextplus{6} -\rail@endplus -\rail@end -\rail@begin{9}{} -\rail@term{\hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}}}[] -\rail@plus -\rail@nont{\isa{class}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@cr{3} -\rail@plus -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{target}}[] -\rail@cr{5} -\rail@plus -\rail@bar -\rail@nextbar{6} -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@endbar -\rail@nextplus{7} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextplus{8} -\rail@endplus -\rail@end -\rail@begin{7}{} -\rail@term{\hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}}}[] -\rail@plus -\rail@nont{\isa{typeconstructor}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\isa{class}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@cr{3} -\rail@plus -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{target}}[] -\rail@plus -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@endbar -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextplus{6} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}}}[] -\rail@nont{\isa{target}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}}}[] -\rail@nont{\isa{const}}[] -\rail@nont{\isa{const}}[] -\rail@nont{\isa{target}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}}}[] -\rail@nont{\isa{target}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}}}[] -\rail@nont{\isa{target}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{11}{} -\rail@term{\hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@cr{2} -\rail@bar -\rail@nextbar{3} -\rail@term{\isa{\isakeyword{datatypes}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{4} -\rail@plus -\rail@plus -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@nextplus{6} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@endbar -\rail@cr{8} -\rail@bar -\rail@nextbar{9} -\rail@term{\isa{\isakeyword{functions}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextplus{10} -\rail@endplus -\rail@endbar -\rail@bar -\rail@nextbar{9} -\rail@term{\isa{\isakeyword{file}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@endbar -\rail@end -\rail@begin{6}{} -\rail@term{\hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}}}[] -\rail@cr{2} -\rail@bar -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{modes}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\isa{modedecl}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{5} -\rail@nont{\isa{const}}[] -\rail@end -\rail@begin{4}{\isa{syntax}} -\rail@bar -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextbar{1} -\rail@bar -\rail@term{\isa{\isakeyword{infix}}}[] -\rail@nextbar{2} -\rail@term{\isa{\isakeyword{infixl}}}[] -\rail@nextbar{3} -\rail@term{\isa{\isakeyword{infixr}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@endbar -\rail@end -\rail@begin{6}{\isa{modedecl}} -\rail@bar -\rail@nont{\isa{modes}}[] -\rail@nextbar{1} -\rail@nont{\isa{const}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\isa{modes}}[] -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{and}}}[] -\rail@plus -\rail@nont{\isa{const}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\isa{modes}}[] -\rail@term{\isa{\isakeyword{and}}}[] -\rail@nextplus{5} -\rail@endplus -\rail@endbar -\rail@endbar -\rail@end -\rail@begin{1}{\isa{modes}} -\rail@nont{\isa{mode}}[] -\rail@term{\isa{\isakeyword{as}}}[] -\rail@nont{\isa{const}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.export-code}{\mbox{\isa{\isacommand{export{\isaliteral{5F}{\isacharunderscore}}code}}}} generates code for a given list - of constants in the specified target language(s). If no - serialization instruction is given, only abstract code is generated - internally. - - Constants may be specified by giving them literally, referring to - all executable contants within a certain theory by giving \isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{2E}{\isachardot}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, or referring to \emph{all} executable constants currently - available by giving \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}. - - By default, for each involved theory one corresponding name space - module is generated. Alternativly, a module name may be specified - after the \hyperlink{keyword.module-name}{\mbox{\isa{\isakeyword{module{\isaliteral{5F}{\isacharunderscore}}name}}}} keyword; then \emph{all} code is - placed in this module. - - For \emph{SML}, \emph{OCaml} and \emph{Scala} the file specification - refers to a single file; for \emph{Haskell}, it refers to a whole - directory, where code is generated in multiple files reflecting the - module hierarchy. Omitting the file specification denotes standard - output. - - Serializers take an optional list of arguments in parentheses. For - \emph{SML} and \emph{OCaml}, ``\isa{no{\isaliteral{5F}{\isacharunderscore}}signatures}`` omits - explicit module signatures. - - For \emph{Haskell} a module name prefix may be given using the - ``\isa{{\isaliteral{22}{\isachardoublequote}}root{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}'' argument; ``\isa{string{\isaliteral{5F}{\isacharunderscore}}classes}'' adds a - ``\verb|deriving (Read, Show)|'' clause to each appropriate - datatype declaration. - - \item \hyperlink{attribute.HOL.code}{\mbox{\isa{code}}} explicitly selects (or with option - ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' deselects) a code equation for code generation. - Usually packages introducing code equations provide a reasonable - default setup for selection. Variants \isa{{\isaliteral{22}{\isachardoublequote}}code\ abstype{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}code\ abstract{\isaliteral{22}{\isachardoublequote}}} declare abstract datatype certificates or - code equations on abstract datatype representations respectively. - - \item \hyperlink{command.HOL.code-abort}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}abort}}}} declares constants which are not - required to have a definition by means of code equations; if needed - these are implemented by program abort instead. - - \item \hyperlink{command.HOL.code-datatype}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}datatype}}}} specifies a constructor set - for a logical type. - - \item \hyperlink{command.HOL.print-codesetup}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codesetup}}}} gives an overview on - selected code equations and code generator datatypes. - - \item \hyperlink{attribute.HOL.code-unfold}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}unfold}}} declares (or with option - ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' removes) theorems which during preprocessing - are applied as rewrite rules to any code equation or evaluation - input. - - \item \hyperlink{attribute.HOL.code-post}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}post}}} declares (or with option ``\isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}}'' removes) theorems which are applied as rewrite rules to any - result of an evaluation. - - \item \hyperlink{attribute.HOL.code-abbrev}{\mbox{\isa{code{\isaliteral{5F}{\isacharunderscore}}abbrev}}} declares equations which are - applied as rewrite rules to any result of an evaluation and - symmetrically during preprocessing to any code equation or evaluation - input. - - \item \hyperlink{command.HOL.print-codeproc}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}codeproc}}}} prints the setup of the code - generator preprocessor. - - \item \hyperlink{command.HOL.code-thms}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}thms}}}} prints a list of theorems - representing the corresponding program containing all given - constants after preprocessing. - - \item \hyperlink{command.HOL.code-deps}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}deps}}}} visualizes dependencies of - theorems representing the corresponding program containing all given - constants after preprocessing. - - \item \hyperlink{command.HOL.code-const}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}const}}}} associates a list of constants - with target-specific serializations; omitting a serialization - deletes an existing serialization. - - \item \hyperlink{command.HOL.code-type}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}type}}}} associates a list of type - constructors with target-specific serializations; omitting a - serialization deletes an existing serialization. - - \item \hyperlink{command.HOL.code-class}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}class}}}} associates a list of classes - with target-specific class names; omitting a serialization deletes - an existing serialization. This applies only to \emph{Haskell}. - - \item \hyperlink{command.HOL.code-instance}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}instance}}}} declares a list of type - constructor / class instance relations as ``already present'' for a - given target. Omitting a ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' deletes an existing - ``already present'' declaration. This applies only to - \emph{Haskell}. - - \item \hyperlink{command.HOL.code-reserved}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reserved}}}} declares a list of names as - reserved for a given target, preventing it to be shadowed by any - generated code. - - \item \hyperlink{command.HOL.code-monad}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}monad}}}} provides an auxiliary mechanism - to generate monadic code for Haskell. - - \item \hyperlink{command.HOL.code-include}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}include}}}} adds arbitrary named content - (``include'') to generated code. A ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' as last argument - will remove an already added ``include''. - - \item \hyperlink{command.HOL.code-modulename}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}modulename}}}} declares aliasings from one - module name onto another. - - \item \hyperlink{command.HOL.code-reflect}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}reflect}}}} without a ``\isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}}'' - argument compiles code into the system runtime environment and - modifies the code generator setup that future invocations of system - runtime code generation referring to one of the ``\isa{{\isaliteral{22}{\isachardoublequote}}datatypes{\isaliteral{22}{\isachardoublequote}}}'' or ``\isa{{\isaliteral{22}{\isachardoublequote}}functions{\isaliteral{22}{\isachardoublequote}}}'' entities use these - precompiled entities. With a ``\isa{{\isaliteral{22}{\isachardoublequote}}file{\isaliteral{22}{\isachardoublequote}}}'' argument, the - corresponding code is generated into that specified file without - modifying the code generator setup. - - \item \hyperlink{command.HOL.code-pred}{\mbox{\isa{\isacommand{code{\isaliteral{5F}{\isacharunderscore}}pred}}}} creates code equations for a - predicate given a set of introduction rules. Optional mode - annotations determine which arguments are supposed to be input or - output. If alternative introduction rules are declared, one must - prove a corresponding elimination rule. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Definition by specification \label{sec:hol-specification}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{HOL}{command}{specification}\hypertarget{command.HOL.specification}{\hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{HOL}{command}{ax\_specification}\hypertarget{command.HOL.ax-specification}{\hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{6}{} -\rail@bar -\rail@term{\hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\isa{decl}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@cr{3} -\rail@plus -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{5} -\rail@endplus -\rail@end -\rail@begin{2}{\isa{decl}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{overloaded}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}decls\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} sets up a - goal stating the existence of terms with the properties specified to - hold for the constants given in \isa{decls}. After finishing the - proof, the theory will be augmented with definitions for the given - constants, as well as with theorems stating the properties for these - constants. - - \item \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}decls\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} sets up - a goal stating the existence of terms with the properties specified - to hold for the constants given in \isa{decls}. After finishing - the proof, the theory will be augmented with axioms expressing the - properties given in the first place. - - \item \isa{decl} declares a constant to be defined by the - specification given. The definition for the constant \isa{c} is - bound to the name \isa{c{\isaliteral{5F}{\isacharunderscore}}def} unless a theorem name is given in - the declaration. Overloaded constants should be declared as such. - - \end{description} - - Whether to use \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}} or \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}} is to some extent a matter of style. \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}} introduces no new axioms, and so by - construction cannot introduce inconsistencies, whereas \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}} does introduce axioms, but only after the - user has explicitly proven it to be safe. A practical issue must be - considered, though: After introducing two constants with the same - properties using \hyperlink{command.HOL.specification}{\mbox{\isa{\isacommand{specification}}}}, one can prove - that the two constants are, in fact, equal. If this might be a - problem, one should use \hyperlink{command.HOL.ax-specification}{\mbox{\isa{\isacommand{ax{\isaliteral{5F}{\isacharunderscore}}specification}}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Inner_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1866 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Inner{\isaliteral{5F}{\isacharunderscore}}Syntax}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Inner{\isaliteral{5F}{\isacharunderscore}}Syntax\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Inner syntax --- the term language \label{ch:inner-syntax}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The inner syntax of Isabelle provides concrete notation for - the main entities of the logical framework, notably \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms with types and type classes. Applications may either - extend existing syntactic categories by additional notation, or - define new sub-languages that are linked to the standard term - language via some explicit markers. For example \verb|FOO|~\isa{{\isaliteral{22}{\isachardoublequote}}foo{\isaliteral{22}{\isachardoublequote}}} could embed the syntax corresponding for some - user-defined nonterminal \isa{{\isaliteral{22}{\isachardoublequote}}foo{\isaliteral{22}{\isachardoublequote}}} --- within the bounds of the - given lexical syntax of Isabelle/Pure. - - The most basic way to specify concrete syntax for logical entities - works via mixfix annotations (\secref{sec:mixfix}), which may be - usually given as part of the original declaration or via explicit - notation commands later on (\secref{sec:notation}). This already - covers many needs of concrete syntax without having to understand - the full complexity of inner syntax layers. - - Further details of the syntax engine involves the classical - distinction of lexical language versus context-free grammar (see - \secref{sec:pure-syntax}), and various mechanisms for \emph{syntax - transformations} (see \secref{sec:syntax-transformations}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Printing logical entities% -} -\isamarkuptrue% -% -\isamarkupsubsection{Diagnostic commands \label{sec:print-diag}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{typ}\hypertarget{command.typ}{\hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{term}\hypertarget{command.term}{\hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{prop}\hypertarget{command.prop}{\hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{thm}\hypertarget{command.thm}{\hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{prf}\hypertarget{command.prf}{\hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{full\_prf}\hypertarget{command.full-prf}{\hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isaliteral{5F}{\isacharunderscore}}prf}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{pr}\hypertarget{command.pr}{\hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - These diagnostic commands assist interactive development by printing - internal logical entities in a human-readable fashion. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isaliteral{5F}{\isacharunderscore}}prf}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{modes}\hypertarget{syntax.modes}{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}~\isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} reads and prints a type expression - according to the current context. - - \item \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ s{\isaliteral{22}{\isachardoublequote}}} uses type-inference to - determine the most general way to make \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} conform to sort - \isa{{\isaliteral{22}{\isachardoublequote}}s{\isaliteral{22}{\isachardoublequote}}}. For concrete \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} this checks if the type - belongs to that sort. Dummy type parameters ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}'' - (underscore) are assigned to fresh type variables with most general - sorts, according the the principles of type-inference. - - \item \hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}~\isa{t} and \hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} - read, type-check and print terms or propositions according to the - current theory or proof context; the inferred type of \isa{t} is - output as well. Note that these commands are also useful in - inspecting the current environment of term abbreviations. - - \item \hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} retrieves - theorems from the current theory or proof context. Note that any - attributes included in the theorem specifications are applied to a - temporary context derived from the current theory or proof; the - result is discarded, i.e.\ attributes involved in \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} do not have any permanent effect. - - \item \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}} displays the (compact) proof term of the - current proof state (if present), or of the given theorems. Note - that this requires proof terms to be switched on for the current - object logic (see the ``Proof terms'' section of the Isabelle - reference manual for information on how to do this). - - \item \hyperlink{command.full-prf}{\mbox{\isa{\isacommand{full{\isaliteral{5F}{\isacharunderscore}}prf}}}} is like \hyperlink{command.prf}{\mbox{\isa{\isacommand{prf}}}}, but displays - the full proof term, i.e.\ also displays information omitted in the - compact proof term, which is denoted by ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' placeholders - there. - - \item \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}goals{\isaliteral{22}{\isachardoublequote}}} prints the current proof state - (if present), including current facts and goals. The optional limit - arguments affect the number of goals to be displayed, which is - initially 10. Omitting limit value leaves the current setting - unchanged. - - \end{description} - - All of the diagnostic commands above admit a list of \isa{modes} - to be specified, which is appended to the current print mode; see - also \secref{sec:print-modes}. Thus the output behavior may be - modified according particular print mode features. For example, - \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}latex\ xsymbols{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} would print the current - proof state with mathematical symbols and special characters - represented in {\LaTeX} source, according to the Isabelle style - \cite{isabelle-sys}. - - Note that antiquotations (cf.\ \secref{sec:antiq}) provide a more - systematic way to include formal items into the printed text - document.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Details of printed content% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{rcll} - \indexdef{}{attribute}{show\_types}\hypertarget{attribute.show-types}{\hyperlink{attribute.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_sorts}\hypertarget{attribute.show-sorts}{\hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_consts}\hypertarget{attribute.show-consts}{\hyperlink{attribute.show-consts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}consts}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_abbrevs}\hypertarget{attribute.show-abbrevs}{\hyperlink{attribute.show-abbrevs}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}} & : & \isa{attribute} & default \isa{true} \\ - \indexdef{}{attribute}{show\_brackets}\hypertarget{attribute.show-brackets}{\hyperlink{attribute.show-brackets}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}brackets}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{names\_long}\hypertarget{attribute.names-long}{\hyperlink{attribute.names-long}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}long}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{names\_short}\hypertarget{attribute.names-short}{\hyperlink{attribute.names-short}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}short}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{names\_unique}\hypertarget{attribute.names-unique}{\hyperlink{attribute.names-unique}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}unique}}}} & : & \isa{attribute} & default \isa{true} \\ - \indexdef{}{attribute}{eta\_contract}\hypertarget{attribute.eta-contract}{\hyperlink{attribute.eta-contract}{\mbox{\isa{eta{\isaliteral{5F}{\isacharunderscore}}contract}}}} & : & \isa{attribute} & default \isa{true} \\ - \indexdef{}{attribute}{goals\_limit}\hypertarget{attribute.goals-limit}{\hyperlink{attribute.goals-limit}{\mbox{\isa{goals{\isaliteral{5F}{\isacharunderscore}}limit}}}} & : & \isa{attribute} & default \isa{{\isadigit{1}}{\isadigit{0}}} \\ - \indexdef{}{attribute}{show\_main\_goal}\hypertarget{attribute.show-main-goal}{\hyperlink{attribute.show-main-goal}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}main{\isaliteral{5F}{\isacharunderscore}}goal}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_hyps}\hypertarget{attribute.show-hyps}{\hyperlink{attribute.show-hyps}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}hyps}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_tags}\hypertarget{attribute.show-tags}{\hyperlink{attribute.show-tags}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}tags}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{show\_question\_marks}\hypertarget{attribute.show-question-marks}{\hyperlink{attribute.show-question-marks}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}question{\isaliteral{5F}{\isacharunderscore}}marks}}}} & : & \isa{attribute} & default \isa{true} \\ - \end{tabular} - \medskip - - These configuration options control the detail of information that - is displayed for types, terms, theorems, goals etc. See also - \secref{sec:config}. - - \begin{description} - - \item \hyperlink{attribute.show-types}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}types}}} and \hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}} control - printing of type constraints for term variables, and sort - constraints for type variables. By default, neither of these are - shown in output. If \hyperlink{attribute.show-sorts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}sorts}}} is enabled, types are - always shown as well. - - Note that displaying types and sorts may explain why a polymorphic - inference rule fails to resolve with some goal, or why a rewrite - rule does not apply as expected. - - \item \hyperlink{attribute.show-consts}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}consts}}} controls printing of types of - constants when displaying a goal state. - - Note that the output can be enormous, because polymorphic constants - often occur at several different type instances. - - \item \hyperlink{attribute.show-abbrevs}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}abbrevs}}} controls folding of constant - abbreviations. - - \item \hyperlink{attribute.show-brackets}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}brackets}}} controls bracketing in pretty - printed output. If enabled, all sub-expressions of the pretty - printing tree will be parenthesized, even if this produces malformed - term syntax! This crude way of showing the internal structure of - pretty printed entities may occasionally help to diagnose problems - with operator priorities, for example. - - \item \hyperlink{attribute.names-long}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}long}}}, \hyperlink{attribute.names-short}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}short}}}, and - \hyperlink{attribute.names-unique}{\mbox{\isa{names{\isaliteral{5F}{\isacharunderscore}}unique}}} control the way of printing fully - qualified internal names in external form. See also - \secref{sec:antiq} for the document antiquotation options of the - same names. - - \item \hyperlink{attribute.eta-contract}{\mbox{\isa{eta{\isaliteral{5F}{\isacharunderscore}}contract}}} controls \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-contracted - printing of terms. - - The \isa{{\isaliteral{5C3C6574613E}{\isasymeta}}}-contraction law asserts \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ f\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ f{\isaliteral{22}{\isachardoublequote}}}, - provided \isa{x} is not free in \isa{f}. It asserts - \emph{extensionality} of functions: \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ g{\isaliteral{22}{\isachardoublequote}}} if \isa{{\isaliteral{22}{\isachardoublequote}}f\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ g\ x{\isaliteral{22}{\isachardoublequote}}} for all \isa{x}. Higher-order unification frequently puts - terms into a fully \isa{{\isaliteral{5C3C6574613E}{\isasymeta}}}-expanded form. For example, if \isa{F} has type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} then its expanded form is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}h{\isaliteral{2E}{\isachardot}}\ F\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ h\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. - - Enabling \hyperlink{attribute.eta-contract}{\mbox{\isa{eta{\isaliteral{5F}{\isacharunderscore}}contract}}} makes Isabelle perform \isa{{\isaliteral{5C3C6574613E}{\isasymeta}}}-contractions before printing, so that \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}h{\isaliteral{2E}{\isachardot}}\ F\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ h\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - appears simply as \isa{F}. - - Note that the distinction between a term and its \isa{{\isaliteral{5C3C6574613E}{\isasymeta}}}-expanded - form occasionally matters. While higher-order resolution and - rewriting operate modulo \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-conversion, some other tools - might look at terms more discretely. - - \item \hyperlink{attribute.goals-limit}{\mbox{\isa{goals{\isaliteral{5F}{\isacharunderscore}}limit}}} controls the maximum number of - subgoals to be shown in goal output. - - \item \hyperlink{attribute.show-main-goal}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}main{\isaliteral{5F}{\isacharunderscore}}goal}}} controls whether the main result - to be proven should be displayed. This information might be - relevant for schematic goals, to inspect the current claim that has - been synthesized so far. - - \item \hyperlink{attribute.show-hyps}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}hyps}}} controls printing of implicit - hypotheses of local facts. Normally, only those hypotheses are - displayed that are \emph{not} covered by the assumptions of the - current context: this situation indicates a fault in some tool being - used. - - By enabling \hyperlink{attribute.show-hyps}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}hyps}}}, output of \emph{all} hypotheses - can be enforced, which is occasionally useful for diagnostic - purposes. - - \item \hyperlink{attribute.show-tags}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}tags}}} controls printing of extra annotations - within theorems, such as internal position information, or the case - names being attached by the attribute \hyperlink{attribute.case-names}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}names}}}. - - Note that the \hyperlink{attribute.tagged}{\mbox{\isa{tagged}}} and \hyperlink{attribute.untagged}{\mbox{\isa{untagged}}} - attributes provide low-level access to the collection of tags - associated with a theorem. - - \item \hyperlink{attribute.show-question-marks}{\mbox{\isa{show{\isaliteral{5F}{\isacharunderscore}}question{\isaliteral{5F}{\isacharunderscore}}marks}}} controls printing of question - marks for schematic variables, such as \isa{{\isaliteral{3F}{\isacharquery}}x}. Only the leading - question mark is affected, the remaining text is unchanged - (including proper markup for schematic variables that might be - relevant for user interfaces). - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Alternative print modes \label{sec:print-modes}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{print\_mode\_value}\verb|print_mode_value: unit -> string list| \\ - \indexdef{}{ML}{Print\_Mode.with\_modes}\verb|Print_Mode.with_modes: string list -> ('a -> 'b) -> 'a -> 'b| \\ - \end{mldecls} - - The \emph{print mode} facility allows to modify various operations - for printing. Commands like \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}, \hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}, - \hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}} (see \secref{sec:print-diag}) take additional print - modes as optional argument. The underlying ML operations are as - follows. - - \begin{description} - - \item \verb|print_mode_value ()| yields the list of currently - active print mode names. This should be understood as symbolic - representation of certain individual features for printing (with - precedence from left to right). - - \item \verb|Print_Mode.with_modes|~\isa{{\isaliteral{22}{\isachardoublequote}}modes\ f\ x{\isaliteral{22}{\isachardoublequote}}} evaluates - \isa{{\isaliteral{22}{\isachardoublequote}}f\ x{\isaliteral{22}{\isachardoublequote}}} in an execution context where the print mode is - prepended by the given \isa{{\isaliteral{22}{\isachardoublequote}}modes{\isaliteral{22}{\isachardoublequote}}}. This provides a thread-safe - way to augment print modes. It is also monotonic in the set of mode - names: it retains the default print mode that certain - user-interfaces might have installed for their proper functioning! - - \end{description} - - \begin{warn} - The old global reference \verb|print_mode| should never be used - directly in applications. Its main reason for being publicly - accessible is to support historic versions of Proof~General. - \end{warn} - - \medskip The pretty printer for inner syntax maintains alternative - mixfix productions for any print mode name invented by the user, say - in commands like \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}} or \hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}. - Mode names can be arbitrary, but the following ones have a specific - meaning by convention: - - \begin{itemize} - - \item \verb|""| (the empty string): default mode; - implicitly active as last element in the list of modes. - - \item \verb|input|: dummy print mode that is never active; may - be used to specify notation that is only available for input. - - \item \verb|internal| dummy print mode that is never active; - used internally in Isabelle/Pure. - - \item \verb|xsymbols|: enable proper mathematical symbols - instead of ASCII art.\footnote{This traditional mode name stems from - the ``X-Symbol'' package for old versions Proof~General with XEmacs, - although that package has been superseded by Unicode in recent - years.} - - \item \verb|HTML|: additional mode that is active in HTML - presentation of Isabelle theory sources; allows to provide - alternative output notation. - - \item \verb|latex|: additional mode that is active in {\LaTeX} - document preparation of Isabelle theory sources; allows to provide - alternative output notation. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Printing limits% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{mldecls} - \indexdef{}{ML}{Pretty.margin\_default}\verb|Pretty.margin_default: int Unsynchronized.ref| \\ - \indexdef{}{ML}{print\_depth}\verb|print_depth: int -> unit| \\ - \end{mldecls} - - These ML functions set limits for pretty printed text. - - \begin{description} - - \item \verb|Pretty.margin_default| indicates the global default for - the right margin of the built-in pretty printer, with initial value - 76. Note that user-interfaces typically control margins - automatically when resizing windows, or even bypass the formatting - engine of Isabelle/ML altogether and do it within the front end via - Isabelle/Scala. - - \item \verb|print_depth|~\isa{n} limits the printing depth of the - ML toplevel pretty printer; the precise effect depends on the ML - compiler and run-time system. Typically \isa{n} should be less - than 10. Bigger values such as 100--1000 are useful for debugging. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Mixfix annotations \label{sec:mixfix}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Mixfix annotations specify concrete \emph{inner syntax} of - Isabelle types and terms. Locally fixed parameters in toplevel - theorem statements, locale and class specifications also admit - mixfix annotations in a fairly uniform manner. A mixfix annotation - describes describes the concrete syntax, the translation to abstract - syntax, and the pretty printing. Special case annotations provide a - simple means of specifying infix operators and binders. - - Isabelle mixfix syntax is inspired by {\OBJ} \cite{OBJ}. It allows - to specify any context-free priority grammar, which is more general - than the fixity declarations of ML and Prolog. - - \begin{railoutput} -\rail@begin{1}{\indexdef{}{syntax}{mixfix}\hypertarget{syntax.mixfix}{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{mfix}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{struct\_mixfix}\hypertarget{syntax.struct-mixfix}{\hyperlink{syntax.struct-mixfix}{\mbox{\isa{struct{\isaliteral{5F}{\isacharunderscore}}mixfix}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nont{\isa{mfix}}[] -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{structure}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{7}{\isa{mfix}} -\rail@bar -\rail@nont{\hyperlink{syntax.template}{\mbox{\isa{template}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{prios}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@nextbar{2} -\rail@bar -\rail@term{\isa{\isakeyword{infix}}}[] -\rail@nextbar{3} -\rail@term{\isa{\isakeyword{infixl}}}[] -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{infixr}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.template}{\mbox{\isa{template}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{5} -\rail@term{\isa{\isakeyword{binder}}}[] -\rail@nont{\hyperlink{syntax.template}{\mbox{\isa{template}}}}[] -\rail@bar -\rail@nextbar{6} -\rail@nont{\isa{prios}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\isa{template}} -\rail@nont{\isa{string}}[] -\rail@end -\rail@begin{2}{\isa{prios}} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@end -\end{railoutput} - - - The string given as \isa{template} may include literal text, - spacing, blocks, and arguments (denoted by ``\isa{{\isaliteral{5F}{\isacharunderscore}}}''); the - special symbol ``\verb|\|'' (printed as ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E6465783E}{\isasymindex}}{\isaliteral{22}{\isachardoublequote}}}'') - represents an index argument that specifies an implicit structure - reference (see also \secref{sec:locale}). Infix and binder - declarations provide common abbreviations for particular mixfix - declarations. So in practice, mixfix templates mostly degenerate to - literal text for concrete syntax, such as ``\verb|++|'' for - an infix symbol.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The general mixfix form% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In full generality, mixfix declarations work as follows. - Suppose a constant \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} is annotated by - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mixfix\ {\isaliteral{5B}{\isacharbrackleft}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}\ p{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, where \isa{{\isaliteral{22}{\isachardoublequote}}mixfix{\isaliteral{22}{\isachardoublequote}}} is a string - \isa{{\isaliteral{22}{\isachardoublequote}}d\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}\ {\isaliteral{5F}{\isacharunderscore}}\ d\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5F}{\isacharunderscore}}\ d\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} consisting of delimiters that surround - argument positions as indicated by underscores. - - Altogether this determines a production for a context-free priority - grammar, where for each argument \isa{{\isaliteral{22}{\isachardoublequote}}i{\isaliteral{22}{\isachardoublequote}}} the syntactic category - is determined by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} (with priority \isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}}), and the - result category is determined from \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} (with priority \isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}). Priority specifications are optional, with default 0 for - arguments and 1000 for the result.\footnote{Omitting priorities is - prone to syntactic ambiguities unless the delimiter tokens determine - fully bracketed notation, as in \isa{{\isaliteral{22}{\isachardoublequote}}if\ {\isaliteral{5F}{\isacharunderscore}}\ then\ {\isaliteral{5F}{\isacharunderscore}}\ else\ {\isaliteral{5F}{\isacharunderscore}}\ fi{\isaliteral{22}{\isachardoublequote}}}.} - - Since \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} may be again a function type, the constant - type scheme may have more argument positions than the mixfix - pattern. Printing a nested application \isa{{\isaliteral{22}{\isachardoublequote}}c\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} for - \isa{{\isaliteral{22}{\isachardoublequote}}m\ {\isaliteral{3E}{\isachargreater}}\ n{\isaliteral{22}{\isachardoublequote}}} works by attaching concrete notation only to the - innermost part, essentially by printing \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} - instead. If a term has fewer arguments than specified in the mixfix - template, the concrete syntax is ignored. - - \medskip A mixfix template may also contain additional directives - for pretty printing, notably spaces, blocks, and breaks. The - general template format is a sequence over any of the following - entities. - - \begin{description} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}d{\isaliteral{22}{\isachardoublequote}}} is a delimiter, namely a non-empty sequence of - characters other than the following special characters: - - \smallskip - \begin{tabular}{ll} - \verb|'| & single quote \\ - \verb|_| & underscore \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E6465783E}{\isasymindex}}{\isaliteral{22}{\isachardoublequote}}} & index symbol \\ - \verb|(| & open parenthesis \\ - \verb|)| & close parenthesis \\ - \verb|/| & slash \\ - \end{tabular} - \medskip - - \item \verb|'| escapes the special meaning of these - meta-characters, producing a literal version of the following - character, unless that is a blank. - - A single quote followed by a blank separates delimiters, without - affecting printing, but input tokens may have additional white space - here. - - \item \verb|_| is an argument position, which stands for a - certain syntactic category in the underlying grammar. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E6465783E}{\isasymindex}}{\isaliteral{22}{\isachardoublequote}}} is an indexed argument position; this is the place - where implicit structure arguments can be attached. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}s{\isaliteral{22}{\isachardoublequote}}} is a non-empty sequence of spaces for printing. - This and the following specifications do not affect parsing at all. - - \item \verb|(|\isa{n} opens a pretty printing block. The - optional number specifies how much indentation to add when a line - break occurs within the block. If the parenthesis is not followed - by digits, the indentation defaults to 0. A block specified via - \verb|(00| is unbreakable. - - \item \verb|)| closes a pretty printing block. - - \item \verb|//| forces a line break. - - \item \verb|/|\isa{s} allows a line break. Here \isa{s} - stands for the string of spaces (zero or more) right after the - slash. These spaces are printed if the break is \emph{not} taken. - - \end{description} - - The general idea of pretty printing with blocks and breaks is also - described in \cite{paulson-ml2}; it goes back to \cite{Oppen:1980}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Infixes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Infix operators are specified by convenient short forms that - abbreviate general mixfix annotations as follows: - - \begin{center} - \begin{tabular}{lll} - - \verb|(|\indexdef{}{keyword}{infix}\hypertarget{keyword.infix}{\hyperlink{keyword.infix}{\mbox{\isa{\isakeyword{infix}}}}}~\verb|"|\isa{sy}\verb|"| \isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|)| - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6D617073746F3E}{\isasymmapsto}}{\isaliteral{22}{\isachardoublequote}}} & - \verb|("(_ |\isa{sy}\verb|/ _)" [|\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}\verb|, |\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}\verb|]|\isa{{\isaliteral{22}{\isachardoublequote}}\ p{\isaliteral{22}{\isachardoublequote}}}\verb|)| \\ - \verb|(|\indexdef{}{keyword}{infixl}\hypertarget{keyword.infixl}{\hyperlink{keyword.infixl}{\mbox{\isa{\isakeyword{infixl}}}}}~\verb|"|\isa{sy}\verb|"| \isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|)| - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6D617073746F3E}{\isasymmapsto}}{\isaliteral{22}{\isachardoublequote}}} & - \verb|("(_ |\isa{sy}\verb|/ _)" [|\isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|, |\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}\verb|]|\isa{{\isaliteral{22}{\isachardoublequote}}\ p{\isaliteral{22}{\isachardoublequote}}}\verb|)| \\ - \verb|(|\indexdef{}{keyword}{infixr}\hypertarget{keyword.infixr}{\hyperlink{keyword.infixr}{\mbox{\isa{\isakeyword{infixr}}}}}~\verb|"|\isa{sy}\verb|"| \isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|)| - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6D617073746F3E}{\isasymmapsto}}{\isaliteral{22}{\isachardoublequote}}} & - \verb|("(_ |\isa{sy}\verb|/ _)" [|\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}\verb|, |\isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|]|\isa{{\isaliteral{22}{\isachardoublequote}}\ p{\isaliteral{22}{\isachardoublequote}}}\verb|)| \\ - - \end{tabular} - \end{center} - - The mixfix template \verb|"(_ |\isa{sy}\verb|/ _)"| - specifies two argument positions; the delimiter is preceded by a - space and followed by a space or line break; the entire phrase is a - pretty printing block. - - The alternative notation \verb|op|~\isa{sy} is introduced - in addition. Thus any infix operator may be written in prefix form - (as in ML), independently of the number of arguments in the term.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Binders% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A \emph{binder} is a variable-binding construct such as a - quantifier. The idea to formalize \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{22}{\isachardoublequote}}} as \isa{{\isaliteral{22}{\isachardoublequote}}All\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} for \isa{{\isaliteral{22}{\isachardoublequote}}All\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequote}}} already goes back - to \cite{church40}. Isabelle declarations of certain higher-order - operators may be annotated with \indexdef{}{keyword}{binder}\hypertarget{keyword.binder}{\hyperlink{keyword.binder}{\mbox{\isa{\isakeyword{binder}}}}} annotations - as follows: - - \begin{center} - \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequote}}}\verb|"|\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}}\verb|" (|\hyperlink{keyword.binder}{\mbox{\isa{\isakeyword{binder}}}}\verb| "|\isa{{\isaliteral{22}{\isachardoublequote}}sy{\isaliteral{22}{\isachardoublequote}}}\verb|" [|\isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|] |\isa{{\isaliteral{22}{\isachardoublequote}}q{\isaliteral{22}{\isachardoublequote}}}\verb|)| - \end{center} - - This introduces concrete binder syntax \isa{{\isaliteral{22}{\isachardoublequote}}sy\ x{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{22}{\isachardoublequote}}}, where - \isa{x} is a bound variable of type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, the body \isa{b} has type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} and the whole term has type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}}. - The optional integer \isa{p} specifies the syntactic priority of - the body; the default is \isa{{\isaliteral{22}{\isachardoublequote}}q{\isaliteral{22}{\isachardoublequote}}}, which is also the priority of - the whole construct. - - Internally, the binder syntax is expanded to something like this: - \begin{center} - \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{5F}{\isacharunderscore}}binder\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequote}}}\verb|"|\isa{{\isaliteral{22}{\isachardoublequote}}idts\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}}\verb|" ("(3|\isa{sy}\verb|_./ _)" [0, |\isa{{\isaliteral{22}{\isachardoublequote}}p{\isaliteral{22}{\isachardoublequote}}}\verb|] |\isa{{\isaliteral{22}{\isachardoublequote}}q{\isaliteral{22}{\isachardoublequote}}}\verb|)| - \end{center} - - Here \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}} is the nonterminal symbol for a list of - identifiers with optional type constraints (see also - \secref{sec:pure-grammar}). The mixfix template \verb|"(3|\isa{sy}\verb|_./ _)"| defines argument positions - for the bound identifiers and the body, separated by a dot with - optional line break; the entire phrase is a pretty printing block of - indentation level 3. Note that there is no extra space after \isa{{\isaliteral{22}{\isachardoublequote}}sy{\isaliteral{22}{\isachardoublequote}}}, so it needs to be included user specification if the binder - syntax ends with a token that may be continued by an identifier - token at the start of \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}. - - Furthermore, a syntax translation to transforms \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{5F}{\isacharunderscore}}binder\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ b{\isaliteral{22}{\isachardoublequote}}} into iterated application \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. - This works in both directions, for parsing and printing.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Explicit notation \label{sec:notation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{type\_notation}\hypertarget{command.type-notation}{\hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}notation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{no\_type\_notation}\hypertarget{command.no-type-notation}{\hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}type{\isaliteral{5F}{\isacharunderscore}}notation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{notation}\hypertarget{command.notation}{\hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{no\_notation}\hypertarget{command.no-notation}{\hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}notation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{write}\hypertarget{command.write}{\hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Commands that introduce new logical entities (terms or types) - usually allow to provide mixfix annotations on the spot, which is - convenient for default notation. Nonetheless, the syntax may be - modified later on by declarations for explicit notation. This - allows to add or delete mixfix annotations for of existing logical - entities within the current context. - - \begin{railoutput} -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}notation}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}type{\isaliteral{5F}{\isacharunderscore}}notation}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[] -\rail@endbar -\rail@cr{3} -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@nextplus{4} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}notation}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[] -\rail@endbar -\rail@cr{3} -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nont{\hyperlink{syntax.struct-mixfix}{\mbox{\isa{struct{\isaliteral{5F}{\isacharunderscore}}mixfix}}}}[] -\rail@nextplus{4} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nont{\hyperlink{syntax.struct-mixfix}{\mbox{\isa{struct{\isaliteral{5F}{\isacharunderscore}}mixfix}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}notation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} associates mixfix - syntax with an existing type constructor. The arity of the - constructor is retrieved from the context. - - \item \hyperlink{command.no-type-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}type{\isaliteral{5F}{\isacharunderscore}}notation}}}} is similar to \hyperlink{command.type-notation}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}notation}}}}, but removes the specified syntax annotation from - the present context. - - \item \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} associates mixfix - syntax with an existing constant or fixed variable. The type - declaration of the given entity is retrieved from the context. - - \item \hyperlink{command.no-notation}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}notation}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}, - but removes the specified syntax annotation from the present - context. - - \item \hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}} is similar to \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}, but - works within an Isar proof body. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{The Pure syntax \label{sec:pure-syntax}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Lexical matters \label{sec:inner-lex}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The inner lexical syntax vaguely resembles the outer one - (\secref{sec:outer-lex}), but some details are different. There are - two main categories of inner syntax tokens: - - \begin{enumerate} - - \item \emph{delimiters} --- the literal tokens occurring in - productions of the given priority grammar (cf.\ - \secref{sec:priority-grammar}); - - \item \emph{named tokens} --- various categories of identifiers etc. - - \end{enumerate} - - Delimiters override named tokens and may thus render certain - identifiers inaccessible. Sometimes the logical context admits - alternative ways to refer to the same entity, potentially via - qualified names. - - \medskip The categories for named tokens are defined once and for - all as follows, reusing some categories of the outer token syntax - (\secref{sec:outer-lex}). - - \begin{center} - \begin{supertabular}{rcl} - \indexdef{inner}{syntax}{id}\hypertarget{syntax.inner.id}{\hyperlink{syntax.inner.id}{\mbox{\isa{id}}}} & = & \indexref{}{syntax}{ident}\hyperlink{syntax.ident}{\mbox{\isa{ident}}} \\ - \indexdef{inner}{syntax}{longid}\hypertarget{syntax.inner.longid}{\hyperlink{syntax.inner.longid}{\mbox{\isa{longid}}}} & = & \indexref{}{syntax}{longident}\hyperlink{syntax.longident}{\mbox{\isa{longident}}} \\ - \indexdef{inner}{syntax}{var}\hypertarget{syntax.inner.var}{\hyperlink{syntax.inner.var}{\mbox{\isa{var}}}} & = & \indexref{}{syntax}{var}\hyperlink{syntax.var}{\mbox{\isa{var}}} \\ - \indexdef{inner}{syntax}{tid}\hypertarget{syntax.inner.tid}{\hyperlink{syntax.inner.tid}{\mbox{\isa{tid}}}} & = & \indexref{}{syntax}{typefree}\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}} \\ - \indexdef{inner}{syntax}{tvar}\hypertarget{syntax.inner.tvar}{\hyperlink{syntax.inner.tvar}{\mbox{\isa{tvar}}}} & = & \indexref{}{syntax}{typevar}\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}} \\ - \indexdef{inner}{syntax}{num\_token}\hypertarget{syntax.inner.num-token}{\hyperlink{syntax.inner.num-token}{\mbox{\isa{num{\isaliteral{5F}{\isacharunderscore}}token}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\ - \indexdef{inner}{syntax}{float\_token}\hypertarget{syntax.inner.float-token}{\hyperlink{syntax.inner.float-token}{\mbox{\isa{float{\isaliteral{5F}{\isacharunderscore}}token}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\ - \indexdef{inner}{syntax}{xnum\_token}\hypertarget{syntax.inner.xnum-token}{\hyperlink{syntax.inner.xnum-token}{\mbox{\isa{xnum{\isaliteral{5F}{\isacharunderscore}}token}}}} & = & \verb|#|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|#-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\ - - \indexdef{inner}{syntax}{str\_token}\hypertarget{syntax.inner.str-token}{\hyperlink{syntax.inner.str-token}{\mbox{\isa{str{\isaliteral{5F}{\isacharunderscore}}token}}}} & = & \verb|''| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|''| \\ - \end{supertabular} - \end{center} - - The token categories \hyperlink{syntax.inner.num-token}{\mbox{\isa{num{\isaliteral{5F}{\isacharunderscore}}token}}}, \hyperlink{syntax.inner.float-token}{\mbox{\isa{float{\isaliteral{5F}{\isacharunderscore}}token}}}, \hyperlink{syntax.inner.xnum-token}{\mbox{\isa{xnum{\isaliteral{5F}{\isacharunderscore}}token}}}, and \hyperlink{syntax.inner.str-token}{\mbox{\isa{str{\isaliteral{5F}{\isacharunderscore}}token}}} are not used in Pure. Object-logics may implement numerals - and string constants by adding appropriate syntax declarations, - together with some translation functions (e.g.\ see Isabelle/HOL). - - The derived categories \indexdef{inner}{syntax}{num\_const}\hypertarget{syntax.inner.num-const}{\hyperlink{syntax.inner.num-const}{\mbox{\isa{num{\isaliteral{5F}{\isacharunderscore}}const}}}}, \indexdef{inner}{syntax}{float\_const}\hypertarget{syntax.inner.float-const}{\hyperlink{syntax.inner.float-const}{\mbox{\isa{float{\isaliteral{5F}{\isacharunderscore}}const}}}}, and \indexdef{inner}{syntax}{num\_const}\hypertarget{syntax.inner.num-const}{\hyperlink{syntax.inner.num-const}{\mbox{\isa{num{\isaliteral{5F}{\isacharunderscore}}const}}}} provide - robust access to the respective tokens: the syntax tree holds a - syntactic constant instead of a free variable.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Priority grammars \label{sec:priority-grammar}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A context-free grammar consists of a set of \emph{terminal - symbols}, a set of \emph{nonterminal symbols} and a set of - \emph{productions}. Productions have the form \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C67616D6D613E}{\isasymgamma}}{\isaliteral{22}{\isachardoublequote}}}, - where \isa{A} is a nonterminal and \isa{{\isaliteral{5C3C67616D6D613E}{\isasymgamma}}} is a string of - terminals and nonterminals. One designated nonterminal is called - the \emph{root symbol}. The language defined by the grammar - consists of all strings of terminals that can be derived from the - root symbol by applying productions as rewrite rules. - - The standard Isabelle parser for inner syntax uses a \emph{priority - grammar}. Each nonterminal is decorated by an integer priority: - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup p\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. In a derivation, \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup p\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} may be rewritten - using a production \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup q\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C67616D6D613E}{\isasymgamma}}{\isaliteral{22}{\isachardoublequote}}} only if \isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{5C3C6C653E}{\isasymle}}\ q{\isaliteral{22}{\isachardoublequote}}}. Any - priority grammar can be translated into a normal context-free - grammar by introducing new nonterminals and productions. - - \medskip Formally, a set of context free productions \isa{G} - induces a derivation relation \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\isaliteral{5C3C5E7375623E}{}\isactrlsub G{\isaliteral{22}{\isachardoublequote}}} as follows. Let \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}} and \isa{{\isaliteral{5C3C626574613E}{\isasymbeta}}} denote strings of terminal or nonterminal symbols. - Then \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup p\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\isaliteral{5C3C5E7375623E}{}\isactrlsub G\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C67616D6D613E}{\isasymgamma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}} holds if and only if \isa{G} - contains some production \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup q\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C67616D6D613E}{\isasymgamma}}{\isaliteral{22}{\isachardoublequote}}} for \isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{5C3C6C653E}{\isasymle}}\ q{\isaliteral{22}{\isachardoublequote}}}. - - \medskip The following grammar for arithmetic expressions - demonstrates how binding power and associativity of operators can be - enforced by priorities. - - \begin{center} - \begin{tabular}{rclr} - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \verb|(| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|)| \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \verb|0| \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|+| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|*| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \verb|-| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \end{center} - The choice of priorities determines that \verb|-| binds - tighter than \verb|*|, which binds tighter than \verb|+|. Furthermore \verb|+| associates to the left and - \verb|*| to the right. - - \medskip For clarity, grammars obey these conventions: - \begin{itemize} - - \item All priorities must lie between 0 and 1000. - - \item Priority 0 on the right-hand side and priority 1000 on the - left-hand side may be omitted. - - \item The production \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup p\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}} is written as \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{28}{\isacharparenleft}}p{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, i.e.\ the priority of the left-hand side actually appears in - a column on the far right. - - \item Alternatives are separated by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}}. - - \item Repetition is indicated by dots \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} in an informal - but obvious way. - - \end{itemize} - - Using these conventions, the example grammar specification above - takes the form: - \begin{center} - \begin{tabular}{rclc} - \isa{A} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}} & \verb|(| \isa{A} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|0| & \qquad\qquad \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{A} \verb|+| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|*| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|-| \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \end{center}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The Pure grammar \label{sec:pure-grammar}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The priority grammar of the \isa{{\isaliteral{22}{\isachardoublequote}}Pure{\isaliteral{22}{\isachardoublequote}}} theory is defined - approximately like this: - - \begin{center} - \begin{supertabular}{rclr} - - \indexdef{inner}{syntax}{any}\hypertarget{syntax.inner.any}{\hyperlink{syntax.inner.any}{\mbox{\isa{any}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}prop\ \ {\isaliteral{7C}{\isacharbar}}\ \ logic{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{prop}\hypertarget{syntax.inner.prop}{\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}} & = & \verb|(| \isa{prop} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{4}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|::| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|==| \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|&&&| \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|==>| \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{2}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|[|\verb,|,\verb|| \isa{prop} \verb|;| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|;| \isa{prop} \verb||\verb,|,\verb|]| \verb|==>| \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6272616B6B3E}{\isasymlbrakk}}{\isaliteral{22}{\isachardoublequote}}} \isa{prop} \verb|;| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|;| \isa{prop} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C726272616B6B3E}{\isasymrbrakk}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}prop\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|!!| \isa{idts} \verb|.| \isa{prop} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} \isa{idts} \verb|.| \isa{prop} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|OFCLASS| \verb|(| \isa{type} \verb|,| \isa{logic} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|SORT_CONSTRAINT| \verb|(| \isa{type} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|TERM| \isa{logic} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|PROP| \isa{aprop} \\\\ - - \indexdef{inner}{syntax}{aprop}\hypertarget{syntax.inner.aprop}{\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}}} & = & \verb|(| \isa{aprop} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ longid\ \ {\isaliteral{7C}{\isacharbar}}\ \ var\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|...| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|CONST| \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|CONST| \isa{{\isaliteral{22}{\isachardoublequote}}longid{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|XCONST| \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|XCONST| \isa{{\isaliteral{22}{\isachardoublequote}}longid{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}logic\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ \ any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{logic}\hypertarget{syntax.inner.logic}{\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}} & = & \verb|(| \isa{logic} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}logic\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{4}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|::| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ longid\ \ {\isaliteral{7C}{\isacharbar}}\ \ var\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|...| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|CONST| \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|CONST| \isa{{\isaliteral{22}{\isachardoublequote}}longid{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|XCONST| \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|XCONST| \isa{{\isaliteral{22}{\isachardoublequote}}longid{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}logic\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ \ any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{9}}{\isadigit{9}}{\isadigit{9}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7374727563743E}{\isasymstruct}}\ index\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|%| \isa{pttrns} \verb|.| \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}} \isa{pttrns} \verb|.| \isa{{\isaliteral{22}{\isachardoublequote}}any\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{3}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|op| \verb|==|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|op| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|op| \verb|&&&| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|op| \verb|==>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|op| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|TYPE| \verb|(| \isa{type} \verb|)| \\\\ - - \indexdef{inner}{syntax}{idt}\hypertarget{syntax.inner.idt}{\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}} & = & \verb|(| \isa{idt} \verb|)|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ id\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{id} \verb|::| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|_| \verb|::| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{index}\hypertarget{syntax.inner.index}{\hyperlink{syntax.inner.index}{\mbox{\isa{index}}}} & = & \verb|\<^bsub>| \isa{{\isaliteral{22}{\isachardoublequote}}logic\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|\<^esub>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C696E6465783E}{\isasymindex}}{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{idts}\hypertarget{syntax.inner.idts}{\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}idt\ \ {\isaliteral{7C}{\isacharbar}}\ \ idt\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ idts{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{pttrn}\hypertarget{syntax.inner.pttrn}{\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}}} & = & \isa{idt} \\\\ - - \indexdef{inner}{syntax}{pttrns}\hypertarget{syntax.inner.pttrns}{\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}pttrn\ \ {\isaliteral{7C}{\isacharbar}}\ \ pttrn\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ pttrns{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{type}\hypertarget{syntax.inner.type}{\hyperlink{syntax.inner.type}{\mbox{\isa{type}}}} & = & \verb|(| \isa{type} \verb|)| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}tid\ \ {\isaliteral{7C}{\isacharbar}}\ \ tvar\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}tid{\isaliteral{22}{\isachardoublequote}}} \verb|::| \isa{{\isaliteral{22}{\isachardoublequote}}sort\ \ {\isaliteral{7C}{\isacharbar}}\ \ tvar\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|::| \isa{{\isaliteral{22}{\isachardoublequote}}sort\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_| \verb|::| \isa{{\isaliteral{22}{\isachardoublequote}}sort{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}type{\isaliteral{5F}{\isacharunderscore}}name\ \ {\isaliteral{7C}{\isacharbar}}\ \ type\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{0}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}\ type{\isaliteral{5F}{\isacharunderscore}}name{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|(| \isa{type} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \isa{type} \verb|)| \isa{type{\isaliteral{5F}{\isacharunderscore}}name} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}type\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|=>| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}type\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isadigit{1}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}{\isaliteral{22}{\isachardoublequote}}} \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|[| \isa{type} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \isa{type} \verb|]| \verb|=>| \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|[| \isa{type} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \isa{type} \verb|]| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}{\isaliteral{22}{\isachardoublequote}}} \isa{type} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{inner}{syntax}{type\_name}\hypertarget{syntax.inner.type-name}{\hyperlink{syntax.inner.type-name}{\mbox{\isa{type{\isaliteral{5F}{\isacharunderscore}}name}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ longid{\isaliteral{22}{\isachardoublequote}}} \\\\ - - \indexdef{inner}{syntax}{sort}\hypertarget{syntax.inner.sort}{\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}}} & = & \hyperlink{syntax.class-name}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}name}}}~\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|{}| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|{| \hyperlink{syntax.class-name}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}name}}} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \hyperlink{syntax.class-name}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}name}}} \verb|}| \\ - \indexdef{inner}{syntax}{class\_name}\hypertarget{syntax.inner.class-name}{\hyperlink{syntax.inner.class-name}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}name}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ longid{\isaliteral{22}{\isachardoublequote}}} \\ - \end{supertabular} - \end{center} - - \medskip Here literal terminals are printed \verb|verbatim|; - see also \secref{sec:inner-lex} for further token categories of the - inner syntax. The meaning of the nonterminals defined by the above - grammar is as follows: - - \begin{description} - - \item \indexref{inner}{syntax}{any}\hyperlink{syntax.inner.any}{\mbox{\isa{any}}} denotes any term. - - \item \indexref{inner}{syntax}{prop}\hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} denotes meta-level propositions, - which are terms of type \isa{prop}. The syntax of such formulae of - the meta-logic is carefully distinguished from usual conventions for - object-logics. In particular, plain \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-term notation is - \emph{not} recognized as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}. - - \item \indexref{inner}{syntax}{aprop}\hyperlink{syntax.inner.aprop}{\mbox{\isa{aprop}}} denotes atomic propositions, which - are embedded into regular \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}} by means of an - explicit \verb|PROP| token. - - Terms of type \isa{prop} with non-constant head, e.g.\ a plain - variable, are printed in this form. Constants that yield type \isa{prop} are expected to provide their own concrete syntax; otherwise - the printed version will appear like \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} and - cannot be parsed again as \hyperlink{syntax.inner.prop}{\mbox{\isa{prop}}}. - - \item \indexref{inner}{syntax}{logic}\hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}} denotes arbitrary terms of a - logical type, excluding type \isa{prop}. This is the main - syntactic category of object-logic entities, covering plain \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-term notation (variables, abstraction, application), plus - anything defined by the user. - - When specifying notation for logical entities, all logical types - (excluding \isa{prop}) are \emph{collapsed} to this single category - of \hyperlink{syntax.inner.logic}{\mbox{\isa{logic}}}. - - \item \indexref{inner}{syntax}{index}\hyperlink{syntax.inner.index}{\mbox{\isa{index}}} denotes an optional index term for - indexed syntax. If omitted, it refers to the first \hyperlink{keyword.structure}{\mbox{\isa{\isakeyword{structure}}}} variable in the context. The special dummy ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C696E6465783E}{\isasymindex}}{\isaliteral{22}{\isachardoublequote}}}'' serves as pattern variable in mixfix annotations that - introduce indexed notation. - - \item \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} denotes identifiers, possibly - constrained by types. - - \item \indexref{inner}{syntax}{idts}\hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}} denotes a sequence of \indexref{inner}{syntax}{idt}\hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}}. This is the most basic category for variables in - iterated binders, such as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}}. - - \item \indexref{inner}{syntax}{pttrn}\hyperlink{syntax.inner.pttrn}{\mbox{\isa{pttrn}}} and \indexref{inner}{syntax}{pttrns}\hyperlink{syntax.inner.pttrns}{\mbox{\isa{pttrns}}} - denote patterns for abstraction, cases bindings etc. In Pure, these - categories start as a merely copy of \hyperlink{syntax.inner.idt}{\mbox{\isa{idt}}} and - \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, respectively. Object-logics may add - additional productions for binding forms. - - \item \indexref{inner}{syntax}{type}\hyperlink{syntax.inner.type}{\mbox{\isa{type}}} denotes types of the meta-logic. - - \item \indexref{inner}{syntax}{sort}\hyperlink{syntax.inner.sort}{\mbox{\isa{sort}}} denotes meta-level sorts. - - \end{description} - - Here are some further explanations of certain syntax features. - - \begin{itemize} - - \item In \hyperlink{syntax.inner.idts}{\mbox{\isa{idts}}}, note that \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ y{\isaliteral{22}{\isachardoublequote}}} is - parsed as \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}nat\ y{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, treating \isa{y} like a type - constructor applied to \isa{nat}. To avoid this interpretation, - write \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}\ y{\isaliteral{22}{\isachardoublequote}}} with explicit parentheses. - - \item Similarly, \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{22}{\isachardoublequote}}} is parsed as \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}nat\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. The correct form is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{22}{\isachardoublequote}}} if \isa{y} is last in the - sequence of identifiers. - - \item Type constraints for terms bind very weakly. For example, - \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3C}{\isacharless}}\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{22}{\isachardoublequote}}} is normally parsed as \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{3C}{\isacharless}}\ y{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{22}{\isachardoublequote}}}, unless \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3C}{\isacharless}}{\isaliteral{22}{\isachardoublequote}}} has a very low priority, in which case the - input is likely to be ambiguous. The correct form is \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3C}{\isacharless}}\ {\isaliteral{28}{\isacharparenleft}}y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}. - - \item Constraints may be either written with two literal colons - ``\verb|::|'' or the double-colon symbol \verb|\|, - which actually looks exactly the same in some {\LaTeX} styles. - - \item Dummy variables (written as underscore) may occur in different - roles. - - \begin{description} - - \item A type ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}'' or ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ sort{\isaliteral{22}{\isachardoublequote}}}'' acts like an - anonymous inference parameter, which is filled-in according to the - most general type produced by the type-checking phase. - - \item A bound ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}'' refers to a vacuous abstraction, where - the body does not refer to the binding introduced here. As in the - term \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ {\isaliteral{5F}{\isacharunderscore}}{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{22}{\isachardoublequote}}}, which is \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}}-equivalent to \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{22}{\isachardoublequote}}}. - - \item A free ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}'' refers to an implicit outer binding. - Higher definitional packages usually allow forms like \isa{{\isaliteral{22}{\isachardoublequote}}f\ x\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequote}}}. - - \item A schematic ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}'' (within a term pattern, see - \secref{sec:term-decls}) refers to an anonymous variable that is - implicitly abstracted over its context of locally bound variables. - For example, this allows pattern matching of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7B}{\isacharbraceleft}}x{\isaliteral{2E}{\isachardot}}\ f\ x\ {\isaliteral{3D}{\isacharequal}}\ g\ x{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} against \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7B}{\isacharbraceleft}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5F}{\isacharunderscore}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, or even \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7B}{\isacharbraceleft}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5F}{\isacharunderscore}}{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} by - using both bound and schematic dummies. - - \end{description} - - \item The three literal dots ``\verb|...|'' may be also - written as ellipsis symbol \verb|\|. In both cases this - refers to a special schematic variable, which is bound in the - context. This special term abbreviation works nicely with - calculational reasoning (\secref{sec:calculation}). - - \item \verb|CONST| ensures that the given identifier is treated - as constant term, and passed through the parse tree in fully - internalized form. This is particularly relevant for translation - rules (\secref{sec:syn-trans}), notably on the RHS. - - \item \verb|XCONST| is similar to \verb|CONST|, but - retains the constant name as given. This is only relevant to - translation rules (\secref{sec:syn-trans}), notably on the LHS. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Inspecting the syntax% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{print\_syntax}\hypertarget{command.print-syntax}{\hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}syntax}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{description} - - \item \hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}syntax}}}} prints the inner syntax of the - current context. The output can be quite large; the most important - sections are explained below. - - \begin{description} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}lexicon{\isaliteral{22}{\isachardoublequote}}} lists the delimiters of the inner token - language; see \secref{sec:inner-lex}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}prods{\isaliteral{22}{\isachardoublequote}}} lists the productions of the underlying - priority grammar; see \secref{sec:priority-grammar}. - - The nonterminal \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7375703E}{}\isactrlsup p\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is rendered in plain text as \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{5B}{\isacharbrackleft}}p{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}; delimiters are quoted. Many productions have an extra - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ name{\isaliteral{22}{\isachardoublequote}}}. These names later become the heads of parse - trees; they also guide the pretty printer. - - Productions without such parse tree names are called \emph{copy - productions}. Their right-hand side must have exactly one - nonterminal symbol (or named token). The parser does not create a - new parse tree node for copy productions, but simply returns the - parse tree of the right-hand symbol. - - If the right-hand side of a copy production consists of a single - nonterminal without any delimiters, then it is called a \emph{chain - production}. Chain productions act as abbreviations: conceptually, - they are removed from the grammar by adding new productions. - Priority information attached to chain productions is ignored; only - the dummy value \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} is displayed. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}print\ modes{\isaliteral{22}{\isachardoublequote}}} lists the alternative print modes - provided by this grammar; see \secref{sec:print-modes}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}parse{\isaliteral{5F}{\isacharunderscore}}rules{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}print{\isaliteral{5F}{\isacharunderscore}}rules{\isaliteral{22}{\isachardoublequote}}} relate to - syntax translations (macros); see \secref{sec:syn-trans}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}parse{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}print{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation{\isaliteral{22}{\isachardoublequote}}} list sets of constants that invoke - translation functions for abstract syntax trees, which are only - required in very special situations; see \secref{sec:tr-funs}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}parse{\isaliteral{5F}{\isacharunderscore}}translation{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}print{\isaliteral{5F}{\isacharunderscore}}translation{\isaliteral{22}{\isachardoublequote}}} - list the sets of constants that invoke regular translation - functions; see \secref{sec:tr-funs}. - - \end{description} - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Ambiguity of parsed expressions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{rcll} - \indexdef{}{attribute}{syntax\_ambiguity\_warning}\hypertarget{attribute.syntax-ambiguity-warning}{\hyperlink{attribute.syntax-ambiguity-warning}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}warning}}}} & : & \isa{attribute} & default \isa{true} \\ - \indexdef{}{attribute}{syntax\_ambiguity\_limit}\hypertarget{attribute.syntax-ambiguity-limit}{\hyperlink{attribute.syntax-ambiguity-limit}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}limit}}}} & : & \isa{attribute} & default \isa{{\isadigit{1}}{\isadigit{0}}} \\ - \end{tabular} - - Depending on the grammar and the given input, parsing may be - ambiguous. Isabelle lets the Earley parser enumerate all possible - parse trees, and then tries to make the best out of the situation. - Terms that cannot be type-checked are filtered out, which often - leads to a unique result in the end. Unlike regular type - reconstruction, which is applied to the whole collection of input - terms simultaneously, the filtering stage only treats each given - term in isolation. Filtering is also not attempted for individual - types or raw ASTs (as required for \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}). - - Certain warning or error messages are printed, depending on the - situation and the given configuration options. Parsing ultimately - fails, if multiple results remain after the filtering phase. - - \begin{description} - - \item \hyperlink{attribute.syntax-ambiguity-warning}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}warning}}} controls output of - explicit warning messages about syntax ambiguity. - - \item \hyperlink{attribute.syntax-ambiguity-limit}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}limit}}} determines the number of - resulting parse trees that are shown as part of the printed message - in case of an ambiguity. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Syntax transformations \label{sec:syntax-transformations}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The inner syntax engine of Isabelle provides separate - mechanisms to transform parse trees either as rewrite systems on - first-order ASTs (\secref{sec:syn-trans}), or ML functions on ASTs - or syntactic \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms (\secref{sec:tr-funs}). This works - both for parsing and printing, as outlined in - \figref{fig:parse-print}. - - \begin{figure}[htbp] - \begin{center} - \begin{tabular}{cl} - string & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & lexer + parser \\ - parse tree & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & parse AST translation \\ - AST & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & AST rewriting (macros) \\ - AST & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & parse translation \\ - --- pre-term --- & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & print translation \\ - AST & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & AST rewriting (macros) \\ - AST & \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F776E3E}{\isasymdown}}{\isaliteral{22}{\isachardoublequote}}} & print AST translation \\ - string & - \end{tabular} - \end{center} - \caption{Parsing and printing with translations}\label{fig:parse-print} - \end{figure} - - These intermediate syntax tree formats eventually lead to a pre-term - with all names and binding scopes resolved, but most type - information still missing. Explicit type constraints might be given by - the user, or implicit position information by the system --- both - need to be passed-through carefully by syntax transformations. - - Pre-terms are further processed by the so-called \emph{check} and - \emph{unckeck} phases that are intertwined with type-inference (see - also \cite{isabelle-implementation}). The latter allows to operate - on higher-order abstract syntax with proper binding and type - information already available. - - As a rule of thumb, anything that manipulates bindings of variables - or constants needs to be implemented as syntax transformation (see - below). Anything else is better done via check/uncheck: a prominent - example application is the \hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}} concept of - Isabelle/Pure.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abstract syntax trees \label{sec:ast}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The ML datatype \verb|Ast.ast| explicitly represents the - intermediate AST format that is used for syntax rewriting - (\secref{sec:syn-trans}). It is defined in ML as follows: - \begin{ttbox} - datatype ast = - Constant of string | - Variable of string | - Appl of ast list - \end{ttbox} - - An AST is either an atom (constant or variable) or a list of (at - least two) subtrees. Occasional diagnostic output of ASTs uses - notation that resembles S-expression of LISP. Constant atoms are - shown as quoted strings, variable atoms as non-quoted strings and - applications as a parenthesized list of subtrees. For example, the - AST - \begin{verbatim} -Ast.Appl - [Ast.Constant "_abs", Ast.Variable "x", Ast.Variable "t"] -\end{verbatim} - is pretty-printed as \verb|("_abs" x t)|. Note that - \verb|()| and \verb|(x)| are excluded as ASTs, because - they have too few subtrees. - - \medskip AST application is merely a pro-forma mechanism to indicate - certain syntactic structures. Thus \verb|(c a b)| could mean - either term application or type application, depending on the - syntactic context. - - Nested application like \verb|(("_abs" x t) u)| is also - possible, but ASTs are definitely first-order: the syntax constant - \verb|"_abs"| does not bind the \verb|x| in any way. - Proper bindings are introduced in later stages of the term syntax, - where \verb|("_abs" x t)| becomes an \verb|Abs| node and - occurrences of \verb|x| in \verb|t| are replaced by bound - variables (represented as de-Bruijn indices).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{AST constants versus variables% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Depending on the situation --- input syntax, output syntax, - translation patterns --- the distinction of atomic asts as \verb|Ast.Constant| versus \verb|Ast.Variable| serves slightly different - purposes. - - Input syntax of a term such as \isa{{\isaliteral{22}{\isachardoublequote}}f\ a\ b\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequote}}} does not yet - indicate the scopes of atomic entities \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{2C}{\isacharcomma}}\ a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{2C}{\isacharcomma}}\ c{\isaliteral{22}{\isachardoublequote}}}: they - could be global constants or local variables, even bound ones - depending on the context of the term. \verb|Ast.Variable| leaves - this choice still open: later syntax layers (or translation - functions) may capture such a variable to determine its role - specifically, to make it a constant, bound variable, free variable - etc. In contrast, syntax translations that introduce already known - constants would rather do it via \verb|Ast.Constant| to prevent - accidental re-interpretation later on. - - Output syntax turns term constants into \verb|Ast.Constant| and - variables (free or schematic) into \verb|Ast.Variable|. This - information is precise when printing fully formal \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms. - - In AST translation patterns (\secref{sec:syn-trans}) the system - guesses from the current theory context which atoms should be - treated as constant versus variable for the matching process. - Sometimes this needs to be indicated more explicitly using \isa{{\isaliteral{22}{\isachardoublequote}}CONST\ c{\isaliteral{22}{\isachardoublequote}}} inside the term language. It is also possible to use - \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} declarations (without mixfix annotation) to - enforce that certain unqualified names are always treated as - constant within the syntax machinery. - - \medskip For ASTs that represent the language of types or sorts, the - situation is much simpler, since the concrete syntax already - distinguishes type variables from type constants (constructors). So - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{27}{\isacharprime}}b{\isaliteral{29}{\isacharparenright}}\ foo{\isaliteral{22}{\isachardoublequote}}} corresponds to an AST application of some - constant for \isa{foo} and variable arguments for \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequote}}} and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{27}{\isacharprime}}b{\isaliteral{22}{\isachardoublequote}}}. Note that the postfix application is merely a feature - of the concrete syntax, while in the AST the constructor occurs in - head position.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Authentic syntax names% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Naming constant entities within ASTs is another delicate - issue. Unqualified names are looked up in the name space tables in - the last stage of parsing, after all translations have been applied. - Since syntax transformations do not know about this later name - resolution yet, there can be surprises in boundary cases. - - \emph{Authentic syntax names} for \verb|Ast.Constant| avoid this - problem: the fully-qualified constant name with a special prefix for - its formal category (\isa{{\isaliteral{22}{\isachardoublequote}}class{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}type{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}const{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}fixed{\isaliteral{22}{\isachardoublequote}}}) represents the information faithfully - within the untyped AST format. Accidental overlap with free or - bound variables is excluded as well. Authentic syntax names work - implicitly in the following situations: - - \begin{itemize} - - \item Input of term constants (or fixed variables) that are - introduced by concrete syntax via \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}}: the - correspondence of a particular grammar production to some known term - entity is preserved. - - \item Input of type constants (constructors) and type classes --- - thanks to explicit syntactic distinction independently on the - context. - - \item Output of term constants, type constants, type classes --- - this information is already available from the internal term to be - printed. - - \end{itemize} - - In other words, syntax transformations that operate on input terms - written as prefix applications are difficult to make robust. - Luckily, this case rarely occurs in practice, because syntax forms - to be translated usually correspond to some bits of concrete - notation.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Raw syntax and translations \label{sec:syn-trans}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{rcll} - \indexdef{}{command}{nonterminal}\hypertarget{command.nonterminal}{\hyperlink{command.nonterminal}{\mbox{\isa{\isacommand{nonterminal}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{syntax}\hypertarget{command.syntax}{\hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{no\_syntax}\hypertarget{command.no-syntax}{\hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}syntax}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{translations}\hypertarget{command.translations}{\hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{no\_translations}\hypertarget{command.no-translations}{\hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}translations}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{syntax\_ast\_trace}\hypertarget{attribute.syntax-ast-trace}{\hyperlink{attribute.syntax-ast-trace}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}trace}}}} & : & \isa{attribute} & default \isa{false} \\ - \indexdef{}{attribute}{syntax\_ast\_stats}\hypertarget{attribute.syntax-ast-stats}{\hyperlink{attribute.syntax-ast-stats}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}stats}}}} & : & \isa{attribute} & default \isa{false} \\ - \end{tabular} - - Unlike mixfix notation for existing formal entities - (\secref{sec:notation}), raw syntax declarations provide full access - to the priority grammar of the inner syntax, without any sanity - checks. This includes additional syntactic categories (via - \hyperlink{command.nonterminal}{\mbox{\isa{\isacommand{nonterminal}}}}) and free-form grammar productions (via - \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}). Additional syntax translations (or macros, via - \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}) are required to turn resulting parse trees - into proper representations of formal entities again. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.nonterminal}{\mbox{\isa{\isacommand{nonterminal}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}syntax}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\isa{constdecl}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{7}{} -\rail@bar -\rail@term{\hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}translations}}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\isa{transpat}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{5C3C72696768746C656674686172706F6F6E733E}{\isasymrightleftharpoons}}}}[] -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{5C3C7269676874686172706F6F6E75703E}{\isasymrightharpoonup}}}}[] -\rail@nextbar{5} -\rail@term{\isa{{\isaliteral{5C3C6C656674686172706F6F6E646F776E3E}{\isasymleftharpoondown}}}}[] -\rail@endbar -\rail@nont{\isa{transpat}}[] -\rail@nextplus{6} -\rail@endplus -\rail@end -\rail@begin{2}{\isa{constdecl}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{mode}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{output}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{\isakeyword{output}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\isa{transpat}} -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.nonterminal}{\mbox{\isa{\isacommand{nonterminal}}}}~\isa{c} declares a type - constructor \isa{c} (without arguments) to act as purely syntactic - type: a nonterminal symbol of the inner syntax. - - \item \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}\ c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} augments the - priority grammar and the pretty printer table for the given print - mode (default \verb|""|). An optional keyword \indexref{}{keyword}{output}\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}} means that only the pretty printer table is affected. - - Following \secref{sec:mixfix}, the mixfix annotation \isa{{\isaliteral{22}{\isachardoublequote}}mx\ {\isaliteral{3D}{\isacharequal}}\ template\ ps\ q{\isaliteral{22}{\isachardoublequote}}} together with type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7369676D613E}{\isasymsigma}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} and - specify a grammar production. The \isa{template} contains - delimiter tokens that surround \isa{{\isaliteral{22}{\isachardoublequote}}n{\isaliteral{22}{\isachardoublequote}}} argument positions - (\verb|_|). The latter correspond to nonterminal symbols - \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} derived from the argument types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} as - follows: - \begin{itemize} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}prop{\isaliteral{22}{\isachardoublequote}}} if \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3D}{\isacharequal}}\ prop{\isaliteral{22}{\isachardoublequote}}} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}logic{\isaliteral{22}{\isachardoublequote}}} if \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{22}{\isachardoublequote}}} for logical type - constructor \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6B617070613E}{\isasymkappa}}\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ prop{\isaliteral{22}{\isachardoublequote}}} - - \item \isa{any} if \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}} for type variables - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{22}{\isachardoublequote}}} if \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub i\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{22}{\isachardoublequote}}} for nonterminal \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6B617070613E}{\isasymkappa}}{\isaliteral{22}{\isachardoublequote}}} - (syntactic type constructor) - - \end{itemize} - - Each \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} is decorated by priority \isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} from the - given list \isa{{\isaliteral{22}{\isachardoublequote}}ps{\isaliteral{22}{\isachardoublequote}}}; misssing priorities default to 0. - - The resulting nonterminal of the production is determined similarly - from type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}, with priority \isa{{\isaliteral{22}{\isachardoublequote}}q{\isaliteral{22}{\isachardoublequote}}} and default 1000. - - \medskip Parsing via this production produces parse trees \isa{{\isaliteral{22}{\isachardoublequote}}t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} for the argument slots. The resulting parse tree is - composed as \isa{{\isaliteral{22}{\isachardoublequote}}c\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, by using the syntax constant \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}} of the syntax declaration. - - Such syntactic constants are invented on the spot, without formal - check wrt.\ existing declarations. It is conventional to use plain - identifiers prefixed by a single underscore (e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}foobar{\isaliteral{22}{\isachardoublequote}}}). Names should be chosen with care, to avoid clashes - with other syntax declarations. - - \medskip The special case of copy production is specified by \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequote}}}\verb|""| (empty string). It means that the - resulting parse tree \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} is copied directly, without any - further decoration. - - \item \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}syntax}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}mode{\isaliteral{29}{\isacharparenright}}\ decls{\isaliteral{22}{\isachardoublequote}}} removes grammar - declarations (and translations) resulting from \isa{decls}, which - are interpreted in the same manner as for \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} above. - - \item \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}~\isa{rules} specifies syntactic - translation rules (i.e.\ macros) as first-order rewrite rules on - ASTs (\secref{sec:ast}). The theory context maintains two - independent lists translation rules: parse rules (\verb|=>| - or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7269676874686172706F6F6E75703E}{\isasymrightharpoonup}}{\isaliteral{22}{\isachardoublequote}}}) and print rules (\verb|<=| or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C656674686172706F6F6E646F776E3E}{\isasymleftharpoondown}}{\isaliteral{22}{\isachardoublequote}}}). - For convenience, both can be specified simultaneously as parse~/ - print rules (\verb|==| or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C72696768746C656674686172706F6F6E733E}{\isasymrightleftharpoons}}{\isaliteral{22}{\isachardoublequote}}}). - - Translation patterns may be prefixed by the syntactic category to be - used for parsing; the default is \isa{logic} which means that - regular term syntax is used. Both sides of the syntax translation - rule undergo parsing and parse AST translations - \secref{sec:tr-funs}, in order to perform some fundamental - normalization like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ b\ {\isaliteral{5C3C6C65616473746F3E}{\isasymleadsto}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}y{\isaliteral{2E}{\isachardot}}\ b{\isaliteral{22}{\isachardoublequote}}}, but other AST - translation rules are \emph{not} applied recursively here. - - When processing AST patterns, the inner syntax lexer runs in a - different mode that allows identifiers to start with underscore. - This accommodates the usual naming convention for auxiliary syntax - constants --- those that do not have a logical counter part --- by - allowing to specify arbitrary AST applications within the term - syntax, independently of the corresponding concrete syntax. - - Atomic ASTs are distinguished as \verb|Ast.Constant| versus \verb|Ast.Variable| as follows: a qualified name or syntax constant - declared via \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}}, or parse tree head of concrete - notation becomes \verb|Ast.Constant|, anything else \verb|Ast.Variable|. Note that \isa{CONST} and \isa{XCONST} within - the term language (\secref{sec:pure-grammar}) allow to enforce - treatment as constants. - - AST rewrite rules \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}lhs{\isaliteral{2C}{\isacharcomma}}\ rhs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} need to obey the following - side-conditions: - - \begin{itemize} - - \item Rules must be left linear: \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}} must not contain - repeated variables.\footnote{The deeper reason for this is that AST - equality is not well-defined: different occurrences of the ``same'' - AST could be decorated differently by accidental type-constraints or - source position information, for example.} - - \item Every variable in \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}} must also occur in \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}}. - - \end{itemize} - - \item \hyperlink{command.no-translations}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}translations}}}}~\isa{rules} removes syntactic - translation rules, which are interpreted in the same manner as for - \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}} above. - - \item \hyperlink{attribute.syntax-ast-trace}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}trace}}} and \hyperlink{attribute.syntax-ast-stats}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}stats}}} control diagnostic output in the AST normalization - process, when translation rules are applied to concrete input or - output. - - \end{description} - - Raw syntax and translations provides a slightly more low-level - access to the grammar and the form of resulting parse trees. It is - often possible to avoid this untyped macro mechanism, and use - type-safe \hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}} or \hyperlink{command.notation}{\mbox{\isa{\isacommand{notation}}}} instead. - Some important situations where \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} and \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}} are really need are as follows: - - \begin{itemize} - - \item Iterated replacement via recursive \hyperlink{command.translations}{\mbox{\isa{\isacommand{translations}}}}. - For example, consider list enumeration \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}a{\isaliteral{2C}{\isacharcomma}}\ b{\isaliteral{2C}{\isacharcomma}}\ c{\isaliteral{2C}{\isacharcomma}}\ d{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} as - defined in theory \isa{List} in Isabelle/HOL. - - \item Change of binding status of variables: anything beyond the - built-in \hyperlink{keyword.binder}{\mbox{\isa{\isakeyword{binder}}}} mixfix annotation requires explicit - syntax translations. For example, consider list filter - comprehension \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}x\ {\isaliteral{5C3C6C6566746172726F773E}{\isasymleftarrow}}\ xs\ {\isaliteral{2E}{\isachardot}}\ P{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} as defined in theory \isa{List} in Isabelle/HOL. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Applying translation rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -As a term is being parsed or printed, an AST is generated as - an intermediate form according to \figref{fig:parse-print}. The AST - is normalized by applying translation rules in the manner of a - first-order term rewriting system. We first examine how a single - rule is applied. - - Let \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} be the abstract syntax tree to be normalized and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}lhs{\isaliteral{2C}{\isacharcomma}}\ rhs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} some translation rule. A subtree \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}} - of \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}} is called \emph{redex} if it is an instance of \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}}; in this case the pattern \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}} is said to match the - object \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}}. A redex matched by \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}} may be - replaced by the corresponding instance of \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}}, thus - \emph{rewriting} the AST \isa{{\isaliteral{22}{\isachardoublequote}}t{\isaliteral{22}{\isachardoublequote}}}. Matching requires some notion - of \emph{place-holders} in rule patterns: \verb|Ast.Variable| serves - this purpose. - - More precisely, the matching of the object \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}} against the - pattern \isa{{\isaliteral{22}{\isachardoublequote}}lhs{\isaliteral{22}{\isachardoublequote}}} is performed as follows: - - \begin{itemize} - - \item Objects of the form \verb|Ast.Variable|~\isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} or \verb|Ast.Constant|~\isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} are matched by pattern \verb|Ast.Constant|~\isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}}. Thus all atomic ASTs in the object are - treated as (potential) constants, and a successful match makes them - actual constants even before name space resolution (see also - \secref{sec:ast}). - - \item Object \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}} is matched by pattern \verb|Ast.Variable|~\isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}}, binding \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} to \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}}. - - \item Object \verb|Ast.Appl|~\isa{{\isaliteral{22}{\isachardoublequote}}us{\isaliteral{22}{\isachardoublequote}}} is matched by \verb|Ast.Appl|~\isa{{\isaliteral{22}{\isachardoublequote}}ts{\isaliteral{22}{\isachardoublequote}}} if \isa{{\isaliteral{22}{\isachardoublequote}}us{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}ts{\isaliteral{22}{\isachardoublequote}}} have the - same length and each corresponding subtree matches. - - \item In every other case, matching fails. - - \end{itemize} - - A successful match yields a substitution that is applied to \isa{{\isaliteral{22}{\isachardoublequote}}rhs{\isaliteral{22}{\isachardoublequote}}}, generating the instance that replaces \isa{{\isaliteral{22}{\isachardoublequote}}u{\isaliteral{22}{\isachardoublequote}}}. - - Normalizing an AST involves repeatedly applying translation rules - until none are applicable. This works yoyo-like: top-down, - bottom-up, top-down, etc. At each subtree position, rules are - chosen in order of appearance in the theory definitions. - - The configuration options \hyperlink{attribute.syntax-ast-trace}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}trace}}} and - \hyperlink{attribute.syntax-ast-stats}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}stats}}} might help to understand this process - and diagnose problems. - - \begin{warn} - If syntax translation rules work incorrectly, the output of - \indexref{}{command}{print\_syntax}\hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}syntax}}}} with its \emph{rules} sections reveals the - actual internal forms of AST pattern, without potentially confusing - concrete syntax. Recall that AST constants appear as quoted strings - and variables without quotes. - \end{warn} - - \begin{warn} - If \indexref{}{attribute}{eta\_contract}\hyperlink{attribute.eta-contract}{\mbox{\isa{eta{\isaliteral{5F}{\isacharunderscore}}contract}}} is set to \isa{{\isaliteral{22}{\isachardoublequote}}true{\isaliteral{22}{\isachardoublequote}}}, terms - will be \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-contracted \emph{before} the AST rewriter sees - them. Thus some abstraction nodes needed for print rules to match - may vanish. For example, \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} would contract - to \isa{{\isaliteral{22}{\isachardoublequote}}Ball\ A\ P{\isaliteral{22}{\isachardoublequote}}} and the standard print rule would fail to - apply. This problem can be avoided by hand-written ML translation - functions (see also \secref{sec:tr-funs}), which is in fact the same - mechanism used in built-in \hyperlink{keyword.binder}{\mbox{\isa{\isakeyword{binder}}}} declarations. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Syntax translation functions \label{sec:tr-funs}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{parse\_ast\_translation}\hypertarget{command.parse-ast-translation}{\hyperlink{command.parse-ast-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{parse\_translation}\hypertarget{command.parse-translation}{\hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}translation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_translation}\hypertarget{command.print-translation}{\hyperlink{command.print-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}translation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{typed\_print\_translation}\hypertarget{command.typed-print-translation}{\hyperlink{command.typed-print-translation}{\mbox{\isa{\isacommand{typed{\isaliteral{5F}{\isacharunderscore}}print{\isaliteral{5F}{\isacharunderscore}}translation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_ast\_translation}\hypertarget{command.print-ast-translation}{\hyperlink{command.print-ast-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{ML antiquotation}{class\_syntax}\hypertarget{ML antiquotation.class-syntax}{\hyperlink{ML antiquotation.class-syntax}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}syntax}}}} & : & \isa{ML{\isaliteral{5F}{\isacharunderscore}}antiquotation} \\ - \indexdef{}{ML antiquotation}{type\_syntax}\hypertarget{ML antiquotation.type-syntax}{\hyperlink{ML antiquotation.type-syntax}{\mbox{\isa{type{\isaliteral{5F}{\isacharunderscore}}syntax}}}} & : & \isa{ML{\isaliteral{5F}{\isacharunderscore}}antiquotation} \\ - \indexdef{}{ML antiquotation}{const\_syntax}\hypertarget{ML antiquotation.const-syntax}{\hyperlink{ML antiquotation.const-syntax}{\mbox{\isa{const{\isaliteral{5F}{\isacharunderscore}}syntax}}}} & : & \isa{ML{\isaliteral{5F}{\isacharunderscore}}antiquotation} \\ - \indexdef{}{ML antiquotation}{syntax\_const}\hypertarget{ML antiquotation.syntax-const}{\hyperlink{ML antiquotation.syntax-const}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}const}}}} & : & \isa{ML{\isaliteral{5F}{\isacharunderscore}}antiquotation} \\ - \end{matharray} - - Syntax translation functions written in ML admit almost arbitrary - manipulations of inner syntax, at the expense of some complexity and - obscurity in the implementation. - - \begin{railoutput} -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{command.parse-ast-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}translation}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.print-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}translation}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.typed-print-translation}{\mbox{\isa{\isacommand{typed{\isaliteral{5F}{\isacharunderscore}}print{\isaliteral{5F}{\isacharunderscore}}translation}}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{command.print-ast-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{advanced}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\rail@begin{4}{} -\rail@bar -\rail@term{\hyperlink{ML antiquotation.class-syntax}{\mbox{\isa{class{\isaliteral{5F}{\isacharunderscore}}syntax}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{ML antiquotation.type-syntax}{\mbox{\isa{type{\isaliteral{5F}{\isacharunderscore}}syntax}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{ML antiquotation.const-syntax}{\mbox{\isa{const{\isaliteral{5F}{\isacharunderscore}}syntax}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{ML antiquotation.syntax-const}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}const}}}}[] -\rail@endbar -\rail@nont{\isa{name}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}translation}}}} etc. declare syntax translation - functions to the theory. Any of these commands have a single - \hyperlink{syntax.text}{\mbox{\isa{text}}} argument that refers to an ML expression of - appropriate type, which are as follows by default: - - \medskip - {\footnotesize - \begin{tabular}{ll} - \hyperlink{command.parse-ast-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}} & : \verb|(string * (Ast.ast list -> Ast.ast)) list| \\ - \hyperlink{command.parse-translation}{\mbox{\isa{\isacommand{parse{\isaliteral{5F}{\isacharunderscore}}translation}}}} & : \verb|(string * (term list -> term)) list| \\ - \hyperlink{command.print-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}translation}}}} & : \verb|(string * (term list -> term)) list| \\ - \hyperlink{command.typed-print-translation}{\mbox{\isa{\isacommand{typed{\isaliteral{5F}{\isacharunderscore}}print{\isaliteral{5F}{\isacharunderscore}}translation}}}} & : \verb|(string * (typ -> term list -> term)) list| \\ - \hyperlink{command.print-ast-translation}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation}}}} & : \verb|(string * (Ast.ast list -> Ast.ast)) list| \\ - \end{tabular}} - \medskip - - The argument list consists of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c{\isaliteral{2C}{\isacharcomma}}\ tr{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} pairs, where \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}} is the syntax name of the formal entity involved, and \isa{{\isaliteral{22}{\isachardoublequote}}tr{\isaliteral{22}{\isachardoublequote}}} a function that translates a syntax form \isa{{\isaliteral{22}{\isachardoublequote}}c\ args{\isaliteral{22}{\isachardoublequote}}} into - \isa{{\isaliteral{22}{\isachardoublequote}}tr\ args{\isaliteral{22}{\isachardoublequote}}}. The ML naming convention for parse translations - is \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{5F}{\isacharunderscore}}tr{\isaliteral{22}{\isachardoublequote}}} and for print translations \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{5F}{\isacharunderscore}}tr{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}}. - - The \indexref{}{command}{print\_syntax}\hyperlink{command.print-syntax}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}syntax}}}} command displays the sets of names - associated with the translation functions of a theory under \isa{{\isaliteral{22}{\isachardoublequote}}parse{\isaliteral{5F}{\isacharunderscore}}ast{\isaliteral{5F}{\isacharunderscore}}translation{\isaliteral{22}{\isachardoublequote}}} etc. - - If the \verb|(|\hyperlink{keyword.advanced}{\mbox{\isa{\isakeyword{advanced}}}}\verb|)| option is - given, the corresponding translation functions depend on the current - theory or proof context as additional argument. This allows to - implement advanced syntax mechanisms, as translations functions may - refer to specific theory declarations or auxiliary proof data. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}class{\isaliteral{5F}{\isacharunderscore}}syntax\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}type{\isaliteral{5F}{\isacharunderscore}}syntax\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}const{\isaliteral{5F}{\isacharunderscore}}syntax\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} inline the authentic syntax name of the - given formal entities into the ML source. This is the - fully-qualified logical name prefixed by a special marker to - indicate its kind: thus different logical name spaces are properly - distinguished within parse trees. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}const{\isaliteral{5F}{\isacharunderscore}}syntax\ c{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} inlines the name \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}} of - the given syntax constant, having checked that it has been declared - via some \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} commands within the theory context. Note - that the usual naming convention makes syntax constants start with - underscore, to reduce the chance of accidental clashes with other - names occurring in parse trees (unqualified constants etc.). - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{The translation strategy% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The different kinds of translation functions are invoked during - the transformations between parse trees, ASTs and syntactic terms - (cf.\ \figref{fig:parse-print}). Whenever a combination of the form - \isa{{\isaliteral{22}{\isachardoublequote}}c\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} is encountered, and a translation function - \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{22}{\isachardoublequote}}} of appropriate kind is declared for \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}}, the - result is produced by evaluation of \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5B}{\isacharbrackleft}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} in ML. - - For AST translations, the arguments \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are ASTs. A - combination has the form \verb|Ast.Constant|~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{22}{\isachardoublequote}}} or \verb|Ast.Appl|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{22}{\isachardoublequote}}}\verb|Ast.Constant|~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}. - For term translations, the arguments are terms and a combination has - the form \verb|Const|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} or \verb|Const|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{24}{\isachardollar}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{24}{\isachardollar}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{24}{\isachardollar}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Terms allow more sophisticated transformations - than ASTs do, typically involving abstractions and bound - variables. \emph{Typed} print translations may even peek at the type - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} of the constant they are invoked on, although that information - may be inaccurate. - - Regardless of whether they act on ASTs or terms, translation - functions called during the parsing process differ from those for - printing in their overall behaviour: - - \begin{description} - - \item [Parse translations] are applied bottom-up. The arguments are - already in translated form. The translations must not fail; - exceptions trigger an error message. There may be at most one - function associated with any syntactic name. - - \item [Print translations] are applied top-down. They are supplied - with arguments that are partly still in internal form. The result - again undergoes translation; therefore a print translation should - not introduce as head the very constant that invoked it. The - function may raise exception \verb|Match| to indicate failure; in - this event it has no effect. Multiple functions associated with - some syntactic name are tried in the order of declaration in the - theory. - - \end{description} - - Only constant atoms --- constructor \verb|Ast.Constant| for ASTs and - \verb|Const| for terms --- can invoke translation functions. This - means that parse translations can only be associated with parse tree - heads of concrete syntax, or syntactic constants introduced via - other translations. For plain identifiers within the term language, - the status of constant versus variable is not yet know during - parsing. This is in contrast to print translations, where constants - are explicitly known from the given term in its fully internal form.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/ML_Tactic.tex --- a/doc-src/IsarRef/Thy/document/ML_Tactic.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,223 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{ML{\isaliteral{5F}{\isacharunderscore}}Tactic}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ ML{\isaliteral{5F}{\isacharunderscore}}Tactic\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{ML tactic expressions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isar Proof methods closely resemble traditional tactics, when used - in unstructured sequences of \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} commands. - Isabelle/Isar provides emulations for all major ML tactics of - classic Isabelle --- mostly for the sake of easy porting of existing - developments, as actual Isar proof texts would demand much less - diversity of proof methods. - - Unlike tactic expressions in ML, Isar proof methods provide proper - concrete syntax for additional arguments, options, modifiers etc. - Thus a typical method text is usually more concise than the - corresponding ML tactic. Furthermore, the Isar versions of classic - Isabelle tactics often cover several variant forms by a single - method with separate options to tune the behavior. For example, - method \hyperlink{method.simp}{\mbox{\isa{simp}}} replaces all of \verb|simp_tac|~/ \verb|asm_simp_tac|~/ \verb|full_simp_tac|~/ \verb|asm_full_simp_tac|, there - is also concrete syntax for augmenting the Simplifier context (the - current ``simpset'') in a convenient way.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Resolution tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Classic Isabelle provides several variant forms of tactics for - single-step rule applications (based on higher-order resolution). - The space of resolution tactics has the following main dimensions. - - \begin{enumerate} - - \item The ``mode'' of resolution: intro, elim, destruct, or forward - (e.g.\ \verb|resolve_tac|, \verb|eresolve_tac|, \verb|dresolve_tac|, - \verb|forward_tac|). - - \item Optional explicit instantiation (e.g.\ \verb|resolve_tac| vs.\ - \verb|res_inst_tac|). - - \item Abbreviations for singleton arguments (e.g.\ \verb|resolve_tac| - vs.\ \verb|rtac|). - - \end{enumerate} - - Basically, the set of Isar tactic emulations \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}}, - \hyperlink{method.erule-tac}{\mbox{\isa{erule{\isaliteral{5F}{\isacharunderscore}}tac}}}, \hyperlink{method.drule-tac}{\mbox{\isa{drule{\isaliteral{5F}{\isacharunderscore}}tac}}}, \hyperlink{method.frule-tac}{\mbox{\isa{frule{\isaliteral{5F}{\isacharunderscore}}tac}}} (see - \secref{sec:tactics}) would be sufficient to cover the four modes, - either with or without instantiation, and either with single or - multiple arguments. Although it is more convenient in most cases to - use the plain \indexref{Pure}{method}{rule}\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method, or any of its - ``improper'' variants \hyperlink{method.erule}{\mbox{\isa{erule}}}, \hyperlink{method.drule}{\mbox{\isa{drule}}}, \hyperlink{method.frule}{\mbox{\isa{frule}}}. Note that explicit goal addressing is only supported by the - actual \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}} version. - - With this in mind, plain resolution tactics correspond to Isar - methods as follows. - - \medskip - \begin{tabular}{lll} - \verb|rtac|~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}rule\ a{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|resolve_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}rule\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|res_inst_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}ctxt\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}\ a\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & - \isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C494E3E}{\isasymIN}}\ a{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \verb|rtac|~\isa{{\isaliteral{22}{\isachardoublequote}}a\ i{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ {\isaliteral{5B}{\isacharbrackleft}}i{\isaliteral{5D}{\isacharbrackright}}\ a{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|resolve_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}\ i{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ {\isaliteral{5B}{\isacharbrackleft}}i{\isaliteral{5D}{\isacharbrackright}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|res_inst_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}ctxt\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{28}{\isacharparenleft}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}\ a\ i{\isaliteral{22}{\isachardoublequote}}} & & - \isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{5F}{\isacharunderscore}}tac\ {\isaliteral{5B}{\isacharbrackleft}}i{\isaliteral{5D}{\isacharbrackright}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C494E3E}{\isasymIN}}\ a{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - Note that explicit goal addressing may be usually avoided by - changing the order of subgoals with \hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}} or \hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}} (see \secref{sec:tactic-commands}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Simplifier tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The main Simplifier tactics \verb|simp_tac| and variants (cf.\ - \cite{isabelle-ref}) are all covered by the \hyperlink{method.simp}{\mbox{\isa{simp}}} and - \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} methods (see \secref{sec:simplifier}). Note that - there is no individual goal addressing available, simplification - acts either on the first goal (\hyperlink{method.simp}{\mbox{\isa{simp}}}) or all goals - (\hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}}). - - \medskip - \begin{tabular}{lll} - \verb|asm_full_simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \hyperlink{method.simp}{\mbox{\isa{simp}}} \\ - \verb|ALLGOALS|~(\verb|asm_full_simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}}) & & \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} \\[0.5ex] - \verb|simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \hyperlink{method.simp}{\mbox{\isa{simp}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|asm_simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \hyperlink{method.simp}{\mbox{\isa{simp}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}simp{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|full_simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \hyperlink{method.simp}{\mbox{\isa{simp}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}use{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|asm_lr_simp_tac|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{40}{\isacharat}}{\isaliteral{7B}{\isacharbraceleft}}simpset{\isaliteral{7D}{\isacharbraceright}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \hyperlink{method.simp}{\mbox{\isa{simp}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}asm{\isaliteral{5F}{\isacharunderscore}}lr{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Classical Reasoner tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Classical Reasoner provides a rather large number of - variations of automated tactics, such as \verb|blast_tac|, \verb|fast_tac|, \verb|clarify_tac| etc. The corresponding Isar methods - usually share the same base name, such as \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.fast}{\mbox{\isa{fast}}}, \hyperlink{method.clarify}{\mbox{\isa{clarify}}} etc.\ (see \secref{sec:classical}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Miscellaneous tactics% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -There are a few additional tactics defined in various theories of - Isabelle/HOL, some of these also in Isabelle/FOL or Isabelle/ZF. - The most common ones of these may be ported to Isar as follows. - - \medskip - \begin{tabular}{lll} - \verb|stac|~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}subst\ a{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|hyp_subst_tac|~\isa{{\isadigit{1}}} & & \isa{hypsubst} \\ - \verb|strip_tac|~\isa{{\isadigit{1}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}intro\ strip{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|split_all_tac|~\isa{{\isadigit{1}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}simp\ {\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}asm{\isaliteral{5F}{\isacharunderscore}}simp{\isaliteral{29}{\isacharparenright}}\ only{\isaliteral{3A}{\isacharcolon}}\ split{\isaliteral{5F}{\isacharunderscore}}tupled{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}simp\ only{\isaliteral{3A}{\isacharcolon}}\ split{\isaliteral{5F}{\isacharunderscore}}tupled{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6C6573733E}{\isasymlless}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}clarify{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Tacticals% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Classic Isabelle provides a huge amount of tacticals for combination - and modification of existing tactics. This has been greatly reduced - in Isar, providing the bare minimum of combinators only: ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2C}{\isacharcomma}}{\isaliteral{22}{\isachardoublequote}}}'' (sequential composition), ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}}'' (alternative - choices), ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' (try), ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}}'' (repeat at least - once). These are usually sufficient in practice; if all fails, - arbitrary ML tactic code may be invoked via the \hyperlink{method.tactic}{\mbox{\isa{tactic}}} - method (see \secref{sec:tactics}). - - \medskip Common ML tacticals may be expressed directly in Isar as - follows: - - \medskip - \begin{tabular}{lll} - \isa{{\isaliteral{22}{\isachardoublequote}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\verb|THEN|~\isa{{\isaliteral{22}{\isachardoublequote}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{{\isaliteral{22}{\isachardoublequote}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\verb|ORELSE|~\isa{{\isaliteral{22}{\isachardoublequote}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{7C}{\isacharbar}}\ meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|TRY|~\isa{tac} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|REPEAT1|~\isa{tac} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth{\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|REPEAT|~\isa{tac} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}meth{\isaliteral{2B}{\isacharplus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|EVERY|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \verb|FIRST|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}meth\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular} - \medskip - - \medskip \verb|CHANGED| (see \cite{isabelle-implementation}) is - usually not required in Isar, since most basic proof methods already - fail unless there is an actual change in the goal state. - Nevertheless, ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' (try) may be used to accept - \emph{unchanged} results as well. - - \medskip \verb|ALLGOALS|, \verb|SOMEGOAL| etc.\ (see - \cite{isabelle-implementation}) are not available in Isar, since - there is no direct goal addressing. Nevertheless, some basic - methods address all goals internally, notably \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} - (see \secref{sec:simplifier}). Also note that \verb|ALLGOALS| can be - often replaced by ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}}'' (repeat at least once), although - this usually has a different operational behavior: subgoals are - solved in a different order. - - \medskip Iterated resolution, such as - \verb|REPEAT (FIRSTGOAL (resolve_tac ...))|, is usually better - expressed using the \hyperlink{method.intro}{\mbox{\isa{intro}}} and \hyperlink{method.elim}{\mbox{\isa{elim}}} methods of - Isar (see \secref{sec:classical}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Misc.tex --- a/doc-src/IsarRef/Thy/document/Misc.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,283 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Misc}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Misc\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Other commands% -} -\isamarkuptrue% -% -\isamarkupsection{Inspecting the context% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{print\_commands}\hypertarget{command.print-commands}{\hyperlink{command.print-commands}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}commands}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_theory}\hypertarget{command.print-theory}{\hyperlink{command.print-theory}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theory}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_methods}\hypertarget{command.print-methods}{\hyperlink{command.print-methods}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}methods}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_attributes}\hypertarget{command.print-attributes}{\hyperlink{command.print-attributes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}attributes}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_theorems}\hypertarget{command.print-theorems}{\hyperlink{command.print-theorems}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theorems}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{find\_theorems}\hypertarget{command.find-theorems}{\hyperlink{command.find-theorems}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}theorems}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{find\_consts}\hypertarget{command.find-consts}{\hyperlink{command.find-consts}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}consts}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{thm\_deps}\hypertarget{command.thm-deps}{\hyperlink{command.thm-deps}{\mbox{\isa{\isacommand{thm{\isaliteral{5F}{\isacharunderscore}}deps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{unused\_thms}\hypertarget{command.unused-thms}{\hyperlink{command.unused-thms}{\mbox{\isa{\isacommand{unused{\isaliteral{5F}{\isacharunderscore}}thms}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_facts}\hypertarget{command.print-facts}{\hyperlink{command.print-facts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}facts}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_binds}\hypertarget{command.print-binds}{\hyperlink{command.print-binds}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}binds}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.print-theory}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theory}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.print-theorems}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theorems}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@endbar -\rail@end -\rail@begin{6}{} -\rail@term{\hyperlink{command.find-theorems}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}theorems}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{2} -\rail@term{\isa{with{\isaliteral{5F}{\isacharunderscore}}dups}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{4} -\rail@plus -\rail@nextplus{5} -\rail@cnont{\isa{thmcriterion}}[] -\rail@endplus -\rail@end -\rail@begin{7}{\isa{thmcriterion}} -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{name}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@term{\isa{intro}}[] -\rail@nextbar{2} -\rail@term{\isa{elim}}[] -\rail@nextbar{3} -\rail@term{\isa{dest}}[] -\rail@nextbar{4} -\rail@term{\isa{solves}}[] -\rail@nextbar{5} -\rail@term{\isa{simp}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{6} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.find-consts}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}consts}}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{constcriterion}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{constcriterion}} -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{name}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@term{\isa{strict}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.thm-deps}{\mbox{\isa{\isacommand{thm{\isaliteral{5F}{\isacharunderscore}}deps}}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.unused-thms}{\mbox{\isa{\isacommand{unused{\isaliteral{5F}{\isacharunderscore}}thms}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{2} -\rail@endplus -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@plus -\rail@nextplus{2} -\rail@cnont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endplus -\rail@endbar -\rail@end -\end{railoutput} - - - These commands print certain parts of the theory and proof context. - Note that there are some further ones available, such as for the set - of rules declared for simplifications. - - \begin{description} - - \item \hyperlink{command.print-commands}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}commands}}}} prints Isabelle's outer theory - syntax, including keywords and command. - - \item \hyperlink{command.print-theory}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theory}}}} prints the main logical content of - the theory context; the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' option indicates extra - verbosity. - - \item \hyperlink{command.print-methods}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}methods}}}} prints all proof methods - available in the current theory context. - - \item \hyperlink{command.print-attributes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}attributes}}}} prints all attributes - available in the current theory context. - - \item \hyperlink{command.print-theorems}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}theorems}}}} prints theorems resulting from the - last command; the ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' option indicates extra verbosity. - - \item \hyperlink{command.find-theorems}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}theorems}}}}~\isa{criteria} retrieves facts - from the theory or proof context matching all of given search - criteria. The criterion \isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ p{\isaliteral{22}{\isachardoublequote}}} selects all theorems - whose fully qualified name matches pattern \isa{p}, which may - contain ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}'' wildcards. The criteria \isa{intro}, - \isa{elim}, and \isa{dest} select theorems that match the - current goal as introduction, elimination or destruction rules, - respectively. The criterion \isa{{\isaliteral{22}{\isachardoublequote}}solves{\isaliteral{22}{\isachardoublequote}}} returns all rules - that would directly solve the current goal. The criterion - \isa{{\isaliteral{22}{\isachardoublequote}}simp{\isaliteral{3A}{\isacharcolon}}\ t{\isaliteral{22}{\isachardoublequote}}} selects all rewrite rules whose left-hand side - matches the given term. The criterion term \isa{t} selects all - theorems that contain the pattern \isa{t} -- as usual, patterns - may contain occurrences of the dummy ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'', schematic - variables, and type constraints. - - Criteria can be preceded by ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' to select theorems that - do \emph{not} match. Note that giving the empty list of criteria - yields \emph{all} currently known facts. An optional limit for the - number of printed facts may be given; the default is 40. By - default, duplicates are removed from the search result. Use - \isa{with{\isaliteral{5F}{\isacharunderscore}}dups} to display duplicates. - - \item \hyperlink{command.find-consts}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}consts}}}}~\isa{criteria} prints all constants - whose type meets all of the given criteria. The criterion \isa{{\isaliteral{22}{\isachardoublequote}}strict{\isaliteral{3A}{\isacharcolon}}\ ty{\isaliteral{22}{\isachardoublequote}}} is met by any type that matches the type pattern - \isa{ty}. Patterns may contain both the dummy type ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' - and sort constraints. The criterion \isa{ty} is similar, but it - also matches against subtypes. The criterion \isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ p{\isaliteral{22}{\isachardoublequote}}} and - the prefix ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' function as described for \hyperlink{command.find-theorems}{\mbox{\isa{\isacommand{find{\isaliteral{5F}{\isacharunderscore}}theorems}}}}. - - \item \hyperlink{command.thm-deps}{\mbox{\isa{\isacommand{thm{\isaliteral{5F}{\isacharunderscore}}deps}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} - visualizes dependencies of facts, using Isabelle's graph browser - tool (see also \cite{isabelle-sys}). - - \item \hyperlink{command.unused-thms}{\mbox{\isa{\isacommand{unused{\isaliteral{5F}{\isacharunderscore}}thms}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub m\ {\isaliteral{2D}{\isacharminus}}\ B\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} - displays all unused theorems in theories \isa{{\isaliteral{22}{\isachardoublequote}}B\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} - or their parents, but not in \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ A\isaliteral{5C3C5E697375623E}{}\isactrlisub m{\isaliteral{22}{\isachardoublequote}}} or their parents. - If \isa{n} is \isa{{\isadigit{0}}}, the end of the range of theories - defaults to the current theory. If no range is specified, - only the unused theorems in the current theory are displayed. - - \item \hyperlink{command.print-facts}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}facts}}}} prints all local facts of the - current context, both named and unnamed ones. - - \item \hyperlink{command.print-binds}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}binds}}}} prints all term abbreviations - present in the context. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{System commands% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{cd}\hypertarget{command.cd}{\hyperlink{command.cd}{\mbox{\isa{\isacommand{cd}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{pwd}\hypertarget{command.pwd}{\hyperlink{command.pwd}{\mbox{\isa{\isacommand{pwd}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{use\_thy}\hypertarget{command.use-thy}{\hyperlink{command.use-thy}{\mbox{\isa{\isacommand{use{\isaliteral{5F}{\isacharunderscore}}thy}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.cd}{\mbox{\isa{\isacommand{cd}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.use-thy}{\mbox{\isa{\isacommand{use{\isaliteral{5F}{\isacharunderscore}}thy}}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.cd}{\mbox{\isa{\isacommand{cd}}}}~\isa{path} changes the current directory - of the Isabelle process. - - \item \hyperlink{command.pwd}{\mbox{\isa{\isacommand{pwd}}}} prints the current working directory. - - \item \hyperlink{command.use-thy}{\mbox{\isa{\isacommand{use{\isaliteral{5F}{\isacharunderscore}}thy}}}}~\isa{A} preload theory \isa{A}. - These system commands are scarcely used when working interactively, - since loading of theories is done automatically as required. - - \end{description} - - %FIXME proper place (!?) - Isabelle file specification may contain path variables (e.g.\ - \verb|$ISABELLE_HOME|) that are expanded accordingly. Note - that \verb|~| abbreviates \verb|$USER_HOME|, and - \verb|~~| abbreviates \verb|$ISABELLE_HOME|. The - general syntax for path specifications follows POSIX conventions.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Outer_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Outer_Syntax.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,737 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Outer{\isaliteral{5F}{\isacharunderscore}}Syntax}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Outer{\isaliteral{5F}{\isacharunderscore}}Syntax\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Outer syntax --- the theory language \label{ch:outer-syntax}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The rather generic framework of Isabelle/Isar syntax emerges from - three main syntactic categories: \emph{commands} of the top-level - Isar engine (covering theory and proof elements), \emph{methods} for - general goal refinements (analogous to traditional ``tactics''), and - \emph{attributes} for operations on facts (within a certain - context). Subsequently we give a reference of basic syntactic - entities underlying Isabelle/Isar syntax in a bottom-up manner. - Concrete theory and proof language elements will be introduced later - on. - - \medskip In order to get started with writing well-formed - Isabelle/Isar documents, the most important aspect to be noted is - the difference of \emph{inner} versus \emph{outer} syntax. Inner - syntax is that of Isabelle types and terms of the logic, while outer - syntax is that of Isabelle/Isar theory sources (specifications and - proofs). As a general rule, inner syntax entities may occur only as - \emph{atomic entities} within outer syntax. For example, the string - \verb|"x + y"| and identifier \verb|z| are legal term - specifications within a theory, while \verb|x + y| without - quotes is not. - - Printed theory documents usually omit quotes to gain readability - (this is a matter of {\LaTeX} macro setup, say via \verb|\isabellestyle|, see also \cite{isabelle-sys}). Experienced - users of Isabelle/Isar may easily reconstruct the lost technical - information, while mere readers need not care about quotes at all. - - \medskip Isabelle/Isar input may contain any number of input - termination characters ``\verb|;|'' (semicolon) to separate - commands explicitly. This is particularly useful in interactive - shell sessions to make clear where the current command is intended - to end. Otherwise, the interpreter loop will continue to issue a - secondary prompt ``\verb|#|'' until an end-of-command is - clearly recognized from the input syntax, e.g.\ encounter of the - next command keyword. - - More advanced interfaces such as Proof~General \cite{proofgeneral} - do not require explicit semicolons, the amount of input text is - determined automatically by inspecting the present content of the - Emacs text buffer. In the printed presentation of Isabelle/Isar - documents semicolons are omitted altogether for readability. - - \begin{warn} - Proof~General requires certain syntax classification tables in - order to achieve properly synchronized interaction with the - Isabelle/Isar process. These tables need to be consistent with - the Isabelle version and particular logic image to be used in a - running session (common object-logics may well change the outer - syntax). The standard setup should work correctly with any of the - ``official'' logic images derived from Isabelle/HOL (including - HOLCF etc.). Users of alternative logics may need to tell - Proof~General explicitly, e.g.\ by giving an option \verb|-k ZF| - (in conjunction with \verb|-l ZF|, to specify the default - logic image). Note that option \verb|-L| does both - of this at the same time. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Lexical matters \label{sec:outer-lex}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The outer lexical syntax consists of three main categories of - syntax tokens: - - \begin{enumerate} - - \item \emph{major keywords} --- the command names that are available - in the present logic session; - - \item \emph{minor keywords} --- additional literal tokens required - by the syntax of commands; - - \item \emph{named tokens} --- various categories of identifiers etc. - - \end{enumerate} - - Major keywords and minor keywords are guaranteed to be disjoint. - This helps user-interfaces to determine the overall structure of a - theory text, without knowing the full details of command syntax. - Internally, there is some additional information about the kind of - major keywords, which approximates the command type (theory command, - proof command etc.). - - Keywords override named tokens. For example, the presence of a - command called \verb|term| inhibits the identifier \verb|term|, but the string \verb|"term"| can be used instead. - By convention, the outer syntax always allows quoted strings in - addition to identifiers, wherever a named entity is expected. - - When tokenizing a given input sequence, the lexer repeatedly takes - the longest prefix of the input that forms a valid token. Spaces, - tabs, newlines and formfeeds between tokens serve as explicit - separators. - - \medskip The categories for named tokens are defined once and for - all as follows. - - \begin{center} - \begin{supertabular}{rcl} - \indexdef{}{syntax}{ident}\hypertarget{syntax.ident}{\hyperlink{syntax.ident}{\mbox{\isa{ident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}letter\ quasiletter\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{syntax}{longident}\hypertarget{syntax.longident}{\hyperlink{syntax.longident}{\mbox{\isa{longident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}ident{\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}}\verb|.|\isa{{\isaliteral{22}{\isachardoublequote}}ident{\isaliteral{29}{\isacharparenright}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{syntax}{symident}\hypertarget{syntax.symident}{\hyperlink{syntax.symident}{\mbox{\isa{symident}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}sym\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{ident}\verb|>| \\ - \indexdef{}{syntax}{nat}\hypertarget{syntax.nat}{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}digit\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{syntax}{float}\hypertarget{syntax.float}{\hyperlink{syntax.float}{\mbox{\isa{float}}}} & = & \indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}}\verb|.|\indexref{}{syntax}{nat}\hyperlink{syntax.nat}{\mbox{\isa{nat}}} \\ - \indexdef{}{syntax}{var}\hypertarget{syntax.var}{\hyperlink{syntax.var}{\mbox{\isa{var}}}} & = & \verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}ident\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{ident}\verb|.|\isa{nat} \\ - \indexdef{}{syntax}{typefree}\hypertarget{syntax.typefree}{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}} & = & \verb|'|\isa{ident} \\ - \indexdef{}{syntax}{typevar}\hypertarget{syntax.typevar}{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}} & = & \verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}typefree\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{typefree}\verb|.|\isa{nat} \\ - \indexdef{}{syntax}{string}\hypertarget{syntax.string}{\hyperlink{syntax.string}{\mbox{\isa{string}}}} & = & \verb|"| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|"| \\ - \indexdef{}{syntax}{altstring}\hypertarget{syntax.altstring}{\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}}} & = & \verb|`| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|`| \\ - \indexdef{}{syntax}{verbatim}\hypertarget{syntax.verbatim}{\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}} & = & \verb|{*| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|*|\verb|}| \\[1ex] - - \isa{letter} & = & \isa{{\isaliteral{22}{\isachardoublequote}}latin\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{latin}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<|\isa{{\isaliteral{22}{\isachardoublequote}}latin\ latin{\isaliteral{22}{\isachardoublequote}}}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ greek\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\<^isub>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\<^isup>| \\ - \isa{quasiletter} & = & \isa{{\isaliteral{22}{\isachardoublequote}}letter\ \ {\isaliteral{7C}{\isacharbar}}\ \ digit\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|'| \\ - \isa{latin} & = & \verb|a|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|z|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|A|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|Z| \\ - \isa{digit} & = & \verb|0|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|9| \\ - \isa{sym} & = & \verb|!|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|#|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|$|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|%|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|&|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|*|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|+|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|-|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|/|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|<|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|=|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|>|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|?|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|@|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|^|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|_|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb||\verb,|,\verb||\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|~| \\ - \isa{greek} & = & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} \\ - & & \verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\|\isa{{\isaliteral{22}{\isachardoublequote}}\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|\| \\ - \end{supertabular} - \end{center} - - A \indexref{}{syntax}{var}\hyperlink{syntax.var}{\mbox{\isa{var}}} or \indexref{}{syntax}{typevar}\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}} describes an unknown, - which is internally a pair of base name and index (ML type \verb|indexname|). These components are either separated by a dot as in - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{2E}{\isachardot}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{7}}{\isaliteral{2E}{\isachardot}}{\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}} or run together as in \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}. The latter form is possible if the base name does not end - with digits. If the index is 0, it may be dropped altogether: - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{2E}{\isachardot}}{\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} all refer to the - same unknown, with basename \isa{{\isaliteral{22}{\isachardoublequote}}x{\isaliteral{22}{\isachardoublequote}}} and index 0. - - The syntax of \indexref{}{syntax}{string}\hyperlink{syntax.string}{\mbox{\isa{string}}} admits any characters, including - newlines; ``\verb|"|'' (double-quote) and ``\verb|\|'' (backslash) need to be escaped by a backslash; arbitrary - character codes may be specified as ``\verb|\|\isa{ddd}'', - with three decimal digits. Alternative strings according to - \indexref{}{syntax}{altstring}\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}} are analogous, using single back-quotes - instead. - - The body of \indexref{}{syntax}{verbatim}\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}} may consist of any text not - containing ``\verb|*|\verb|}|''; this allows - convenient inclusion of quotes without further escapes. There is no - way to escape ``\verb|*|\verb|}|''. If the quoted - text is {\LaTeX} source, one may usually add some blank or comment - to avoid the critical character sequence. - - Source comments take the form \verb|(*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*)| and may be nested, although the user-interface - might prevent this. Note that this form indicates source comments - only, which are stripped after lexical analysis of the input. The - Isar syntax also provides proper \emph{document comments} that are - considered as part of the text (see \secref{sec:comments}). - - Common mathematical symbols such as \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}} are represented in - Isabelle as \verb|\|. There are infinitely many Isabelle - symbols like this, although proper presentation is left to front-end - tools such as {\LaTeX}, Proof~General, or Isabelle/jEdit. A list of - predefined Isabelle symbols that work well with these tools is given - in \appref{app:symbols}. Note that \verb|\| does not belong - to the \isa{letter} category, since it is already used differently - in the Pure term language.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Common syntax entities% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We now introduce several basic syntactic entities, such as names, - terms, and theorem specifications, which are factored out of the - actual Isar language elements to be described later.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Names% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Entity \hyperlink{syntax.name}{\mbox{\isa{name}}} usually refers to any name of types, - constants, theorems etc.\ that are to be \emph{declared} or - \emph{defined} (so qualified identifiers are excluded here). Quoted - strings provide an escape for non-identifier names or those ruled - out by outer syntax keywords (e.g.\ quoted \verb|"let"|). - Already existing objects are usually referenced by \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}. - - \begin{railoutput} -\rail@begin{4}{\indexdef{}{syntax}{name}\hypertarget{syntax.name}{\hyperlink{syntax.name}{\mbox{\isa{name}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.ident}{\mbox{\isa{ident}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.symident}{\mbox{\isa{symident}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{parname}\hypertarget{syntax.parname}{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{nameref}\hypertarget{syntax.nameref}{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.longident}{\mbox{\isa{longident}}}}[] -\rail@endbar -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Numbers% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The outer lexical syntax (\secref{sec:outer-lex}) admits - natural numbers and floating point numbers. These are combined as - \hyperlink{syntax.int}{\mbox{\isa{int}}} and \hyperlink{syntax.real}{\mbox{\isa{real}}} as follows. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{int}\hypertarget{syntax.int}{\hyperlink{syntax.int}{\mbox{\isa{int}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{real}\hypertarget{syntax.real}{\hyperlink{syntax.real}{\mbox{\isa{real}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.float}{\mbox{\isa{float}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.int}{\mbox{\isa{int}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - Note that there is an overlap with the category \hyperlink{syntax.name}{\mbox{\isa{name}}}, - which also includes \hyperlink{syntax.nat}{\mbox{\isa{nat}}}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Comments \label{sec:comments}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Large chunks of plain \hyperlink{syntax.text}{\mbox{\isa{text}}} are usually given - \hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}, i.e.\ enclosed in \verb|{|\verb|*|~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\verb|*|\verb|}|. For convenience, - any of the smaller text units conforming to \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}} are - admitted as well. A marginal \hyperlink{syntax.comment}{\mbox{\isa{comment}}} is of the form - \verb|--|~\hyperlink{syntax.text}{\mbox{\isa{text}}}. Any number of these may occur - within Isabelle/Isar commands. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{text}\hypertarget{syntax.text}{\hyperlink{syntax.text}{\mbox{\isa{text}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.verbatim}{\mbox{\isa{verbatim}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{comment}\hypertarget{syntax.comment}{\hyperlink{syntax.comment}{\mbox{\isa{comment}}}}} -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}{\isaliteral{2D}{\isacharminus}}}}[] -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Type classes, sorts and arities% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Classes are specified by plain names. Sorts have a very simple - inner syntax, which is either a single class name \isa{c} or a - list \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7B}{\isacharbraceleft}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{7D}{\isacharbraceright}}{\isaliteral{22}{\isachardoublequote}}} referring to the - intersection of these classes. The syntax of type arities is given - directly at the outer level. - - \begin{railoutput} -\rail@begin{3}{\indexdef{}{syntax}{classdecl}\hypertarget{syntax.classdecl}{\hyperlink{syntax.classdecl}{\mbox{\isa{classdecl}}}}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@bar -\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{sort}\hypertarget{syntax.sort}{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\rail@begin{3}{\indexdef{}{syntax}{arity}\hypertarget{syntax.arity}{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}} -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Types and terms \label{sec:types-terms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The actual inner Isabelle syntax, that of types and terms of the - logic, is far too sophisticated in order to be modelled explicitly - at the outer theory level. Basically, any such entity has to be - quoted to turn it into a single token (the parsing and type-checking - is performed internally later). For convenience, a slightly more - liberal convention is adopted: quotes may be omitted for any type or - term that is already atomic at the outer level. For example, one - may just write \verb|x| instead of quoted \verb|"x"|. - Note that symbolic identifiers (e.g.\ \verb|++| or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}} are available as well, provided these have not been superseded - by commands or other keywords already (such as \verb|=| or - \verb|+|). - - \begin{railoutput} -\rail@begin{3}{\indexdef{}{syntax}{type}\hypertarget{syntax.type}{\hyperlink{syntax.type}{\mbox{\isa{type}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{term}\hypertarget{syntax.term}{\hyperlink{syntax.term}{\mbox{\isa{term}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.var}{\mbox{\isa{var}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{prop}\hypertarget{syntax.prop}{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@end -\end{railoutput} - - - Positional instantiations are indicated by giving a sequence of - terms, or the placeholder ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' (underscore), which means to - skip a position. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{inst}\hypertarget{syntax.inst}{\hyperlink{syntax.inst}{\mbox{\isa{inst}}}}} -\rail@bar -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{insts}\hypertarget{syntax.insts}{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}} -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.inst}{\mbox{\isa{inst}}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - Type declarations and definitions usually refer to \hyperlink{syntax.typespec}{\mbox{\isa{typespec}}} on the left-hand side. This models basic type constructor - application at the outer syntax level. Note that only plain postfix - notation is available here, but no infixes. - - \begin{railoutput} -\rail@begin{4}{\indexdef{}{syntax}{typespec}\hypertarget{syntax.typespec}{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@nextplus{3} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\rail@begin{6}{\indexdef{}{syntax}{typespec\_sorts}\hypertarget{syntax.typespec-sorts}{\hyperlink{syntax.typespec-sorts}{\mbox{\isa{typespec{\isaliteral{5F}{\isacharunderscore}}sorts}}}}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@bar -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@endbar -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@endbar -\rail@nextplus{5} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Term patterns and declarations \label{sec:term-decls}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Wherever explicit propositions (or term fragments) occur in a - proof text, casual binding of schematic term variables may be given - specified via patterns of the form ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}''. - This works both for \hyperlink{syntax.term}{\mbox{\isa{term}}} and \hyperlink{syntax.prop}{\mbox{\isa{prop}}}. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{term\_pat}\hypertarget{syntax.term-pat}{\hyperlink{syntax.term-pat}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}pat}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@term{\isa{\isakeyword{is}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{prop\_pat}\hypertarget{syntax.prop-pat}{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@term{\isa{\isakeyword{is}}}[] -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - - - \medskip Declarations of local variables \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} and - logical propositions \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} represent different views on - the same principle of introducing a local scope. In practice, one - may usually omit the typing of \hyperlink{syntax.vars}{\mbox{\isa{vars}}} (due to - type-inference), and the naming of propositions (due to implicit - references of current facts). In any case, Isar proof elements - usually admit to introduce multiple such items simultaneously. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{vars}\hypertarget{syntax.vars}{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}} -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\indexdef{}{syntax}{props}\hypertarget{syntax.props}{\hyperlink{syntax.props}{\mbox{\isa{props}}}}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@endplus -\rail@end -\end{railoutput} - - - The treatment of multiple declarations corresponds to the - complementary focus of \hyperlink{syntax.vars}{\mbox{\isa{vars}}} versus \hyperlink{syntax.props}{\mbox{\isa{props}}}. In - ``\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}'' the typing refers to all variables, while - in \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} the naming refers to all propositions - collectively. Isar language elements that refer to \hyperlink{syntax.vars}{\mbox{\isa{vars}}} - or \hyperlink{syntax.props}{\mbox{\isa{props}}} typically admit separate typings or namings via - another level of iteration, with explicit \indexref{}{keyword}{and}\hyperlink{keyword.and}{\mbox{\isa{\isakeyword{and}}}} - separators; e.g.\ see \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}} and \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} in - \secref{sec:proof-context}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Attributes and theorems \label{sec:syn-att}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Attributes have their own ``semi-inner'' syntax, in the sense - that input conforming to \hyperlink{syntax.args}{\mbox{\isa{args}}} below is parsed by the - attribute a second time. The attribute argument specifications may - be any sequence of atomic entities (identifiers, strings etc.), or - properly bracketed argument lists. Below \hyperlink{syntax.atom}{\mbox{\isa{atom}}} refers to - any atomic entity, including any \hyperlink{syntax.keyword}{\mbox{\isa{keyword}}} conforming to - \hyperlink{syntax.symident}{\mbox{\isa{symident}}}. - - \begin{railoutput} -\rail@begin{7}{\indexdef{}{syntax}{atom}\hypertarget{syntax.atom}{\hyperlink{syntax.atom}{\mbox{\isa{atom}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.var}{\mbox{\isa{var}}}}[] -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{5} -\rail@nont{\hyperlink{syntax.float}{\mbox{\isa{float}}}}[] -\rail@nextbar{6} -\rail@nont{\hyperlink{syntax.keyword}{\mbox{\isa{keyword}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{arg}} -\rail@bar -\rail@nont{\hyperlink{syntax.atom}{\mbox{\isa{atom}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{args}\hypertarget{syntax.args}{\hyperlink{syntax.args}{\mbox{\isa{args}}}}} -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{arg}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\indexdef{}{syntax}{attributes}\hypertarget{syntax.attributes}{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@endbar -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@end -\end{railoutput} - - - Theorem specifications come in several flavors: \hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}} - and \hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}} usually refer to axioms, assumptions or - results of goal statements, while \hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}} collects lists of - existing theorems. Existing theorems are given by \hyperlink{syntax.thmref}{\mbox{\isa{thmref}}} - and \hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}, the former requires an actual singleton - result. - - There are three forms of theorem references: - \begin{enumerate} - - \item named facts \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{22}{\isachardoublequote}}}, - - \item selections from named facts \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{28}{\isacharparenleft}}i{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} or \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{28}{\isacharparenleft}}j\ {\isaliteral{2D}{\isacharminus}}\ k{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}, - - \item literal fact propositions using \indexref{}{syntax}{altstring}\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}} syntax - \verb|`|\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}\verb|`| (see also method - \indexref{}{method}{fact}\hyperlink{method.fact}{\mbox{\isa{fact}}}). - - \end{enumerate} - - Any kind of theorem specification may include lists of attributes - both on the left and right hand sides; attributes are applied to any - immediately preceding fact. If names are omitted, the theorems are - not stored within the theorem database of the theory or proof - context, but any given attributes are applied nonetheless. - - An extra pair of brackets around attributes (like ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}simproc\ a{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'') abbreviates a theorem reference involving an - internal dummy fact, which will be ignored later on. So only the - effect of the attribute on the background context will persist. - This form of in-place declarations is particularly useful with - commands like \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} and \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}. - - \begin{railoutput} -\rail@begin{2}{\indexdef{}{syntax}{axmdecl}\hypertarget{syntax.axmdecl}{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{thmdecl}\hypertarget{syntax.thmdecl}{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}} -\rail@nont{\isa{thmbind}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{thmdef}\hypertarget{syntax.thmdef}{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}} -\rail@nont{\isa{thmbind}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@end -\rail@begin{4}{\indexdef{}{syntax}{thmref}\hypertarget{syntax.thmref}{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}} -\rail@bar -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{selection}}[] -\rail@endbar -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.altstring}{\mbox{\isa{altstring}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[] -\rail@endbar -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[] -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{thmrefs}\hypertarget{syntax.thmrefs}{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}} -\rail@plus -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{3}{\isa{thmbind}} -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.attributes}{\mbox{\isa{attributes}}}}[] -\rail@endbar -\rail@end -\rail@begin{4}{\isa{selection}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@plus -\rail@bar -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@bar -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@endbar -\rail@nextplus{3} -\rail@cterm{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Preface.tex --- a/doc-src/IsarRef/Thy/document/Preface.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,106 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Preface}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Preface\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Preface% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \emph{Isabelle} system essentially provides a generic - infrastructure for building deductive systems (programmed in - Standard ML), with a special focus on interactive theorem proving in - higher-order logics. Many years ago, even end-users would refer to - certain ML functions (goal commands, tactics, tacticals etc.) to - pursue their everyday theorem proving tasks. - - In contrast \emph{Isar} provides an interpreted language environment - of its own, which has been specifically tailored for the needs of - theory and proof development. Compared to raw ML, the Isabelle/Isar - top-level provides a more robust and comfortable development - platform, with proper support for theory development graphs, managed - transactions with unlimited undo etc. The Isabelle/Isar version of - the \emph{Proof~General} user interface - \cite{proofgeneral,Aspinall:TACAS:2000} provides a decent front-end - for interactive theory and proof development in this advanced - theorem proving environment, even though it is somewhat biased - towards old-style proof scripts. - - \medskip Apart from the technical advances over bare-bones ML - programming, the main purpose of the Isar language is to provide a - conceptually different view on machine-checked proofs - \cite{Wenzel:1999:TPHOL,Wenzel-PhD}. \emph{Isar} stands for - \emph{Intelligible semi-automated reasoning}. Drawing from both the - traditions of informal mathematical proof texts and high-level - programming languages, Isar offers a versatile environment for - structured formal proof documents. Thus properly written Isar - proofs become accessible to a broader audience than unstructured - tactic scripts (which typically only provide operational information - for the machine). Writing human-readable proof texts certainly - requires some additional efforts by the writer to achieve a good - presentation, both of formal and informal parts of the text. On the - other hand, human-readable formal texts gain some value in their own - right, independently of the mechanic proof-checking process. - - Despite its grand design of structured proof texts, Isar is able to - assimilate the old tactical style as an ``improper'' sub-language. - This provides an easy upgrade path for existing tactic scripts, as - well as some means for interactive experimentation and debugging of - structured proofs. Isabelle/Isar supports a broad range of proof - styles, both readable and unreadable ones. - - \medskip The generic Isabelle/Isar framework (see - \chref{ch:isar-framework}) works reasonably well for any Isabelle - object-logic that conforms to the natural deduction view of the - Isabelle/Pure framework. Specific language elements introduced by - Isabelle/HOL are described in \chref{ch:hol}. Although the main - language elements are already provided by the Isabelle/Pure - framework, examples given in the generic parts will usually refer to - Isabelle/HOL. - - \medskip Isar commands may be either \emph{proper} document - constructors, or \emph{improper commands}. Some proof methods and - attributes introduced later are classified as improper as well. - Improper Isar language elements, which are marked by ``\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}'' in the subsequent chapters; they are often helpful - when developing proof documents, but their use is discouraged for - the final human-readable outcome. Typical examples are diagnostic - commands that print terms or theorems according to the current - context; other commands emulate old-style tactical theorem proving.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Proof.tex --- a/doc-src/IsarRef/Thy/document/Proof.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2107 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Proof}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Proof\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Proofs \label{ch:proofs}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Proof commands perform transitions of Isar/VM machine - configurations, which are block-structured, consisting of a stack of - nodes with three main components: logical proof context, current - facts, and open goals. Isar/VM transitions are typed according to - the following three different modes of operation: - - \begin{description} - - \item \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} means that a new goal has just been - stated that is now to be \emph{proven}; the next command may refine - it by some proof method, and enter a sub-proof to establish the - actual result. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is like a nested theory mode: the - context may be augmented by \emph{stating} additional assumptions, - intermediate results etc. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} is intermediate between \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}: existing facts (i.e.\ - the contents of the special ``\indexref{}{fact}{this}\hyperlink{fact.this}{\mbox{\isa{this}}}'' register) have been - just picked up in order to be used when refining the goal claimed - next. - - \end{description} - - The proof mode indicator may be understood as an instruction to the - writer, telling what kind of operation may be performed next. The - corresponding typings of proof commands restricts the shape of - well-formed proof texts to particular command sequences. So dynamic - arrangements of commands eventually turn out as static texts of a - certain structure. - - \Appref{ap:refcard} gives a simplified grammar of the (extensible) - language emerging that way from the different types of proof - commands. The main ideas of the overall Isar framework are - explained in \chref{ch:isar-framework}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proof structure% -} -\isamarkuptrue% -% -\isamarkupsubsection{Formal notepad% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{notepad}\hypertarget{command.notepad}{\hyperlink{command.notepad}{\mbox{\isa{\isacommand{notepad}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{command.notepad}{\mbox{\isa{\isacommand{notepad}}}}}[] -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.end}{\mbox{\isa{\isacommand{end}}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.notepad}{\mbox{\isa{\isacommand{notepad}}}}~\hyperlink{keyword.begin}{\mbox{\isa{\isakeyword{begin}}}} opens a proof state - without any goal statement. This allows to experiment with Isar, - without producing any persistent result. - - The notepad can be closed by \hyperlink{command.end}{\mbox{\isa{\isacommand{end}}}} or discontinued by - \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}}. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Blocks% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{next}\hypertarget{command.next}{\hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{\{}\hypertarget{command.braceleft}{\hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{\}}\hypertarget{command.braceright}{\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - While Isar is inherently block-structured, opening and closing - blocks is mostly handled rather casually, with little explicit - user-intervention. Any local goal statement automatically opens - \emph{two} internal blocks, which are closed again when concluding - the sub-proof (by \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} etc.). Sections of different - context within a sub-proof may be switched via \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}}, - which is just a single block-close followed by block-open again. - The effect of \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} is to reset the local proof context; - there is no goal focus involved here! - - For slightly more advanced applications, there are explicit block - parentheses as well. These typically achieve a stronger forward - style of reasoning. - - \begin{description} - - \item \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} switches to a fresh block within a - sub-proof, resetting the local context to the initial one. - - \item \hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}} and \hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}} explicitly open and close - blocks. Any current facts pass through ``\hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}'' - unchanged, while ``\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}}'' causes any result to be - \emph{exported} into the enclosing context. Thus fixed variables - are generalized, assumptions discharged, and local definitions - unfolded (cf.\ \secref{sec:proof-context}). There is no difference - of \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} and \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}} in this mode of - forward reasoning --- in contrast to plain backward reasoning with - the result exported at \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} time. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Omitting proofs% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{oops}\hypertarget{command.oops}{\hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - The \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}} command discontinues the current proof - attempt, while considering the partial proof text as properly - processed. This is conceptually quite different from ``faking'' - actual proofs via \indexref{}{command}{sorry}\hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}} (see - \secref{sec:proof-steps}): \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}} does not observe the - proof structure at all, but goes back right to the theory level. - Furthermore, \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}} does not produce any result theorem - --- there is no intended claim to be able to complete the proof - in any way. - - A typical application of \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}} is to explain Isar proofs - \emph{within} the system itself, in conjunction with the document - preparation tools of Isabelle described in \chref{ch:document-prep}. - Thus partial or even wrong proof attempts can be discussed in a - logically sound manner. Note that the Isabelle {\LaTeX} macros can - be easily adapted to print something like ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' instead of - the keyword ``\hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}}''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Statements% -} -\isamarkuptrue% -% -\isamarkupsubsection{Context elements \label{sec:proof-context}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{fix}\hypertarget{command.fix}{\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{assume}\hypertarget{command.assume}{\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{presume}\hypertarget{command.presume}{\hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{def}\hypertarget{command.def}{\hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - The logical proof context consists of fixed variables and - assumptions. The former closely correspond to Skolem constants, or - meta-level universal quantification as provided by the Isabelle/Pure - logical framework. Introducing some \emph{arbitrary, but fixed} - variable via ``\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}'' results in a local value - that may be used in the subsequent proof as any other variable or - constant. Furthermore, any result \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} exported from - the context will be universally closed wrt.\ \isa{x} at the - outermost level: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} (this is expressed in normal - form using Isabelle's meta-variables). - - Similarly, introducing some assumption \isa{{\isaliteral{5C3C6368693E}{\isasymchi}}} has two effects. - On the one hand, a local theorem is created that may be used as a - fact in subsequent proof steps. On the other hand, any result - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6368693E}{\isasymchi}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} exported from the context becomes conditional wrt.\ - the assumption: \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C6368693E}{\isasymchi}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}. Thus, solving an enclosing goal - using such a result would basically introduce a new subgoal stemming - from the assumption. How this situation is handled depends on the - version of assumption command used: while \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} - insists on solving the subgoal by unification with some premise of - the goal, \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}} leaves the subgoal unchanged in order - to be proved later by the user. - - Local definitions, introduced by ``\hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}'', are achieved by combining ``\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}'' with - another version of assumption that causes any hypothetical equation - \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} to be eliminated by the reflexivity rule. Thus, - exporting some result \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} yields \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{5B}{\isacharbrackleft}}t{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}}[] -\rail@plus -\rail@nont{\isa{def}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{5}{\isa{def}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{5C3C65717569763E}{\isasymequiv}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.term-pat}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}pat}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x} introduces a local variable \isa{x} that is \emph{arbitrary, but fixed.} - - \item \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} and \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} introduce a local fact \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} by - assumption. Subsequent results applied to an enclosing goal (e.g.\ - by \indexref{}{command}{show}\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}) are handled as follows: \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} expects to be able to unify with existing premises in the - goal, while \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}} leaves \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} as new subgoals. - - Several lists of assumptions may be given (separated by - \indexref{}{keyword}{and}\hyperlink{keyword.and}{\mbox{\isa{\isakeyword{and}}}}; the resulting list of current facts consists - of all of these concatenated. - - \item \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} introduces a local - (non-polymorphic) definition. In results exported from the context, - \isa{x} is replaced by \isa{t}. Basically, ``\hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}'' abbreviates ``\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}'', with the resulting - hypothetical equation solved by reflexivity. - - The default name for the definitional equation is \isa{x{\isaliteral{5F}{\isacharunderscore}}def}. - Several simultaneous definitions may be given at the same time. - - \end{description} - - The special name \indexref{}{fact}{prems}\hyperlink{fact.prems}{\mbox{\isa{prems}}} refers to all assumptions of the - current context as a list of theorems. This feature should be used - with great care! It is better avoided in final proof texts.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Term abbreviations \label{sec:term-abbrev}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{let}\hypertarget{command.let}{\hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{keyword}{is}\hypertarget{keyword.is}{\hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}}} & : & syntax \\ - \end{matharray} - - Abbreviations may be either bound by explicit \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} statements, or by annotating assumptions or - goal statements with a list of patterns ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}''. In both cases, higher-order matching is invoked to - bind extra-logical term variables, which may be either named - schematic variables of the form \isa{{\isaliteral{3F}{\isacharquery}}x}, or nameless dummies - ``\hyperlink{variable.underscore}{\mbox{\isa{{\isaliteral{5F}{\isacharunderscore}}}}}'' (underscore). Note that in the \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}} - form the patterns occur on the left-hand side, while the \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} patterns are in postfix position. - - Polymorphism of term bindings is handled in Hindley-Milner style, - similar to ML. Type variables referring to local assumptions or - open goal statements are \emph{fixed}, while those of finished - results or bound by \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}} may occur in \emph{arbitrary} - instances later. Even though actual polymorphism should be rarely - used in practice, this mechanism is essential to achieve proper - incremental type-inference, as the user proceeds to build up the - Isar proof text from left to right. - - \medskip Term abbreviations are quite different from local - definitions as introduced via \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} (see - \secref{sec:proof-context}). The latter are visible within the - logic as actual equations, while abbreviations disappear during the - input process just after type checking. Also note that \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} does not support polymorphism. - - \begin{railoutput} -\rail@begin{3}{} -\rail@term{\hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}}[] -\rail@plus -\rail@plus -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - The syntax of \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} patterns follows \hyperlink{syntax.term-pat}{\mbox{\isa{term{\isaliteral{5F}{\isacharunderscore}}pat}}} or - \hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}} (see \secref{sec:term-decls}). - - \begin{description} - - \item \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} binds any - text variables in patterns \isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} by simultaneous - higher-order matching against terms \isa{{\isaliteral{22}{\isachardoublequote}}t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} resembles \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}, but - matches \isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} against the preceding statement. Also - note that \hyperlink{keyword.is}{\mbox{\isa{\isakeyword{is}}}} is not a separate command, but part of - others (such as \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} etc.). - - \end{description} - - Some \emph{implicit} term abbreviations\index{term abbreviations} - for goals and facts are available as well. For any open goal, - \indexref{}{variable}{thesis}\hyperlink{variable.thesis}{\mbox{\isa{thesis}}} refers to its object-level statement, - abstracted over any meta-level parameters (if present). Likewise, - \indexref{}{variable}{this}\hyperlink{variable.this}{\mbox{\isa{this}}} is bound for fact statements resulting from - assumptions or finished goals. In case \hyperlink{variable.this}{\mbox{\isa{this}}} refers to - an object-logic statement that is an application \isa{{\isaliteral{22}{\isachardoublequote}}f\ t{\isaliteral{22}{\isachardoublequote}}}, then - \isa{t} is bound to the special text variable ``\hyperlink{variable.dots}{\mbox{\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}}}'' - (three dots). The canonical application of this convenience are - calculational proofs (see \secref{sec:calculation}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Facts and forward chaining \label{sec:proof-facts}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{note}\hypertarget{command.note}{\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{then}\hypertarget{command.then}{\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{from}\hypertarget{command.from}{\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{with}\hypertarget{command.with}{\hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{using}\hypertarget{command.using}{\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{unfolding}\hypertarget{command.unfolding}{\hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - New facts are established either by assumption or proof of local - statements. Any fact will usually be involved in further proofs, - either as explicit arguments of proof methods, or when forward - chaining towards the next goal via \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} (and variants); - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}} and \hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}} are composite forms - involving \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}. The \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}} elements - augments the collection of used facts \emph{after} a goal has been - stated. Note that the special theorem name \indexref{}{fact}{this}\hyperlink{fact.this}{\mbox{\isa{this}}} refers - to the most recently established facts, but only \emph{before} - issuing a follow-up claim. - - \begin{railoutput} -\rail@begin{3}{} -\rail@term{\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}}[] -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{4}{} -\rail@bar -\rail@term{\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} recalls existing facts - \isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}, binding the result as \isa{a}. Note that - attributes may be involved as well, both on the left and right hand - sides. - - \item \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} indicates forward chaining by the current - facts in order to establish the goal to be claimed next. The - initial proof method invoked to refine that will be offered the - facts to do ``anything appropriate'' (see also - \secref{sec:proof-steps}). For example, method \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} - (see \secref{sec:pure-meth-att}) would typically do an elimination - rather than an introduction. Automatic methods usually insert the - facts into the goal state before operation. This provides a simple - scheme to control relevance of facts in automated proof search. - - \item \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{b} abbreviates ``\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{b}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}''; thus \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} is - equivalent to ``\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this}''. - - \item \hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} abbreviates ``\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ this{\isaliteral{22}{\isachardoublequote}}}''; thus the forward chaining - is from earlier facts together with the current ones. - - \item \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} augments the facts being - currently indicated for use by a subsequent refinement step (such as - \indexref{}{command}{apply}\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} or \indexref{}{command}{proof}\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}). - - \item \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} is structurally - similar to \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}, but unfolds definitional equations - \isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} throughout the goal state and facts. - - \end{description} - - Forward chaining with an empty list of theorems is the same as not - chaining at all. Thus ``\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{nothing}'' has no - effect apart from entering \isa{{\isaliteral{22}{\isachardoublequote}}prove{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} mode, since - \indexref{}{fact}{nothing}\hyperlink{fact.nothing}{\mbox{\isa{nothing}}} is bound to the empty list of theorems. - - Basic proof methods (such as \indexref{Pure}{method}{rule}\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}) expect multiple - facts to be given in their proper order, corresponding to a prefix - of the premises of the rule involved. Note that positions may be - easily skipped using something like \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ a\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ b{\isaliteral{22}{\isachardoublequote}}}, for example. This involves the trivial rule - \isa{{\isaliteral{22}{\isachardoublequote}}PROP\ {\isaliteral{5C3C7073693E}{\isasympsi}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ PROP\ {\isaliteral{5C3C7073693E}{\isasympsi}}{\isaliteral{22}{\isachardoublequote}}}, which is bound in Isabelle/Pure as - ``\indexref{}{fact}{\_}\hyperlink{fact.underscore}{\mbox{\isa{{\isaliteral{5F}{\isacharunderscore}}}}}'' (underscore). - - Automated methods (such as \hyperlink{method.simp}{\mbox{\isa{simp}}} or \hyperlink{method.auto}{\mbox{\isa{auto}}}) just - insert any given facts before their usual operation. Depending on - the kind of procedure involved, the order of facts is less - significant here.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Goals \label{sec:goals}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{lemma}\hypertarget{command.lemma}{\hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{theorem}\hypertarget{command.theorem}{\hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{corollary}\hypertarget{command.corollary}{\hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{schematic\_lemma}\hypertarget{command.schematic-lemma}{\hyperlink{command.schematic-lemma}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}lemma}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{schematic\_theorem}\hypertarget{command.schematic-theorem}{\hyperlink{command.schematic-theorem}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}theorem}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{schematic\_corollary}\hypertarget{command.schematic-corollary}{\hyperlink{command.schematic-corollary}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}corollary}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{have}\hypertarget{command.have}{\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{show}\hypertarget{command.show}{\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{hence}\hypertarget{command.hence}{\hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{thus}\hypertarget{command.thus}{\hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_statement}\hypertarget{command.print-statement}{\hyperlink{command.print-statement}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - From a theory context, proof mode is entered by an initial goal - command such as \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}, \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}, or - \hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}. Within a proof, new claims may be - introduced locally as well; four variants are available here to - indicate whether forward chaining of facts should be performed - initially (via \indexref{}{command}{then}\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}), and whether the final result - is meant to solve some pending goal. - - Goals may consist of multiple statements, resulting in a list of - facts eventually. A pending multi-goal is internally represented as - a meta-level conjunction (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{26}{\isacharampersand}}{\isaliteral{22}{\isachardoublequote}}}), which is usually - split into the corresponding number of sub-goals prior to an initial - method application, via \indexref{}{command}{proof}\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} - (\secref{sec:proof-steps}) or \indexref{}{command}{apply}\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} - (\secref{sec:tactic-commands}). The \indexref{}{method}{induct}\hyperlink{method.induct}{\mbox{\isa{induct}}} method - covered in \secref{sec:cases-induct} acts on multiple claims - simultaneously. - - Claims at the theory level may be either in short or long form. A - short goal merely consists of several simultaneous propositions - (often just one). A long goal includes an explicit context - specification for the subsequent conclusion, involving local - parameters and assumptions. Here the role of each part of the - statement is explicitly marked by separate keywords (see also - \secref{sec:locale}); the local assumptions being introduced here - are available as \indexref{}{fact}{assms}\hyperlink{fact.assms}{\mbox{\isa{assms}}} in the proof. Moreover, there - are two kinds of conclusions: \indexdef{}{element}{shows}\hypertarget{element.shows}{\hyperlink{element.shows}{\mbox{\isa{\isakeyword{shows}}}}} states several - simultaneous propositions (essentially a big conjunction), while - \indexdef{}{element}{obtains}\hypertarget{element.obtains}{\hyperlink{element.obtains}{\mbox{\isa{\isakeyword{obtains}}}}} claims several simultaneous simultaneous - contexts of (essentially a big disjunction of eliminated parameters - and assumptions, cf.\ \secref{sec:obtain}). - - \begin{railoutput} -\rail@begin{6}{} -\rail@bar -\rail@term{\hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.schematic-lemma}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}lemma}}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{command.schematic-theorem}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}theorem}}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{command.schematic-corollary}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}corollary}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nont{\isa{goal}}[] -\rail@nextbar{1} -\rail@nont{\isa{longgoal}}[] -\rail@endbar -\rail@end -\rail@begin{4}{} -\rail@bar -\rail@term{\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}}}[] -\rail@endbar -\rail@nont{\isa{goal}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.print-statement}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.modes}{\mbox{\isa{modes}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{2}{\isa{goal}} -\rail@plus -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{longgoal}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\indexref{}{syntax}{includes}\hyperlink{syntax.includes}{\mbox{\isa{includes}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@endplus -\rail@nont{\isa{conclusion}}[] -\rail@end -\rail@begin{4}{\isa{conclusion}} -\rail@bar -\rail@term{\isa{\isakeyword{shows}}}[] -\rail@nont{\isa{goal}}[] -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{obtains}}}[] -\rail@plus -\rail@bar -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] -\rail@endbar -\rail@nont{\isa{case}}[] -\rail@nextplus{3} -\rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{2}{\isa{case}} -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{\isakeyword{where}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} enters proof mode with - \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} as main goal, eventually resulting in some fact \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} to be put back into the target context. An additional \hyperlink{syntax.context}{\mbox{\isa{context}}} specification may build up an initial proof context for the - subsequent claim; this includes local definitions and syntax as - well, see also \hyperlink{syntax.includes}{\mbox{\isa{includes}}} in \secref{sec:bundle} and - \hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}} in \secref{sec:locale}. - - \item \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} and \hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} are essentially the same as \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}, but the facts are internally marked as - being of a different kind. This discrimination acts like a formal - comment. - - \item \hyperlink{command.schematic-lemma}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}lemma}}}}, \hyperlink{command.schematic-theorem}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}theorem}}}}, - \hyperlink{command.schematic-corollary}{\mbox{\isa{\isacommand{schematic{\isaliteral{5F}{\isacharunderscore}}corollary}}}} are similar to \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}}, - \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}, \hyperlink{command.corollary}{\mbox{\isa{\isacommand{corollary}}}}, respectively but allow - the statement to contain unbound schematic variables. - - Under normal circumstances, an Isar proof text needs to specify - claims explicitly. Schematic goals are more like goals in Prolog, - where certain results are synthesized in the course of reasoning. - With schematic statements, the inherent compositionality of Isar - proofs is lost, which also impacts performance, because proof - checking is forced into sequential mode. - - \item \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} claims a local goal, - eventually resulting in a fact within the current logical context. - This operation is completely independent of any pending sub-goals of - an enclosing goal statements, so \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} may be freely - used for experimental exploration of potential results within a - proof body. - - \item \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} is like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} plus a second stage to refine some pending - sub-goal for each one of the finished result, after having been - exported into the corresponding context (at the head of the - sub-proof of this \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} command). - - To accommodate interactive debugging, resulting rules are printed - before being applied internally. Even more, interactive execution - of \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} predicts potential failure and displays the - resulting error as a warning beforehand. Watch out for the - following message: - - %FIXME proper antiquitation - \begin{ttbox} - Problem! Local statement will fail to solve any pending goal - \end{ttbox} - - \item \hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}} abbreviates ``\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}'', i.e.\ claims a local goal to be proven by forward - chaining the current facts. Note that \hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}} is also - equivalent to ``\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}''. - - \item \hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}} abbreviates ``\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}''. Note that \hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}} is also equivalent to - ``\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}''. - - \item \hyperlink{command.print-statement}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}}}}~\isa{a} prints facts from the - current theory or proof context in long statement form, according to - the syntax for \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}} given above. - - \end{description} - - Any goal statement causes some term abbreviations (such as - \indexref{}{variable}{?thesis}\hyperlink{variable.?thesis}{\mbox{\isa{{\isaliteral{3F}{\isacharquery}}thesis}}}) to be bound automatically, see also - \secref{sec:term-abbrev}. - - The optional case names of \indexref{}{element}{obtains}\hyperlink{element.obtains}{\mbox{\isa{\isakeyword{obtains}}}} have a twofold - meaning: (1) during the of this claim they refer to the the local - context introductions, (2) the resulting rule is annotated - accordingly to support symbolic case splits when used with the - \indexref{}{method}{cases}\hyperlink{method.cases}{\mbox{\isa{cases}}} method (cf.\ \secref{sec:cases-induct}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Refinement steps% -} -\isamarkuptrue% -% -\isamarkupsubsection{Proof method expressions \label{sec:proof-meth}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Proof methods are either basic ones, or expressions composed - of methods via ``\verb|,|'' (sequential composition), - ``\verb||\verb,|,\verb||'' (alternative choices), ``\verb|?|'' - (try), ``\verb|+|'' (repeat at least once), ``\verb|[|\isa{n}\verb|]|'' (restriction to first \isa{n} - sub-goals, with default \isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}). In practice, proof - methods are usually just a comma separated list of \hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}~\hyperlink{syntax.args}{\mbox{\isa{args}}} specifications. Note that parentheses may - be dropped for single method specifications (with no arguments). - - \begin{railoutput} -\rail@begin{5}{\indexdef{}{syntax}{method}\hypertarget{syntax.method}{\hyperlink{syntax.method}{\mbox{\isa{method}}}}} -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{methods}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@end -\rail@begin{4}{\isa{methods}} -\rail@plus -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nont{\hyperlink{syntax.args}{\mbox{\isa{args}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.method}{\mbox{\isa{method}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@bar -\rail@term{\isa{{\isaliteral{2C}{\isacharcomma}}}}[] -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{7C}{\isacharbar}}}}[] -\rail@endbar -\rail@endplus -\rail@end -\end{railoutput} - - - Proper Isar proof methods do \emph{not} admit arbitrary goal - addressing, but refer either to the first sub-goal or all sub-goals - uniformly. The goal restriction operator ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}n{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' - evaluates a method expression within a sandbox consisting of the - first \isa{n} sub-goals (which need to exist). For example, the - method ``\isa{{\isaliteral{22}{\isachardoublequote}}simp{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{5B}{\isacharbrackleft}}{\isadigit{3}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' simplifies the first three - sub-goals, while ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}rule\ foo{\isaliteral{2C}{\isacharcomma}}\ simp{\isaliteral{5F}{\isacharunderscore}}all{\isaliteral{29}{\isacharparenright}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' simplifies all - new goals that emerge from applying rule \isa{{\isaliteral{22}{\isachardoublequote}}foo{\isaliteral{22}{\isachardoublequote}}} to the - originally first one. - - Improper methods, notably tactic emulations, offer a separate - low-level goal addressing scheme as explicit argument to the - individual tactic being involved. Here ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' refers to - all goals, and ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}n{\isaliteral{2D}{\isacharminus}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' to all goals starting from \isa{{\isaliteral{22}{\isachardoublequote}}n{\isaliteral{22}{\isachardoublequote}}}. - - \begin{railoutput} -\rail@begin{4}{\indexdef{}{syntax}{goal\_spec}\hypertarget{syntax.goal-spec}{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{2D}{\isacharminus}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@end -\end{railoutput}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Initial and terminal proof steps \label{sec:proof-steps}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{proof}\hypertarget{command.proof}{\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{qed}\hypertarget{command.qed}{\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{by}\hypertarget{command.by}{\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{..}\hypertarget{command.ddot}{\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{.}\hypertarget{command.dot}{\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{sorry}\hypertarget{command.sorry}{\hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Arbitrary goal refinement via tactics is considered harmful. - Structured proof composition in Isar admits proof methods to be - invoked in two places only. - - \begin{enumerate} - - \item An \emph{initial} refinement step \indexref{}{command}{proof}\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} reduces a newly stated goal to a number - of sub-goals that are to be solved later. Facts are passed to - \isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} for forward chaining, if so indicated by \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} mode. - - \item A \emph{terminal} conclusion step \indexref{}{command}{qed}\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} is intended to solve remaining goals. No facts are - passed to \isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}. - - \end{enumerate} - - The only other (proper) way to affect pending goals in a proof body - is by \indexref{}{command}{show}\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}, which involves an explicit statement of - what is to be solved eventually. Thus we avoid the fundamental - problem of unstructured tactic scripts that consist of numerous - consecutive goal transformations, with invisible effects. - - \medskip As a general rule of thumb for good proof style, initial - proof methods should either solve the goal completely, or constitute - some well-understood reduction to new sub-goals. Arbitrary - automatic proof tools that are prone leave a large number of badly - structured sub-goals are no help in continuing the proof document in - an intelligible manner. - - Unless given explicitly by the user, the default initial method is - \indexref{Pure}{method}{rule}\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} (or its classical variant \indexref{}{method}{rule}\hyperlink{method.rule}{\mbox{\isa{rule}}}), which applies a single standard elimination or introduction - rule according to the topmost symbol involved. There is no separate - default terminal method. Any remaining goals are always solved by - assumption in the very last step. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{method}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{method}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}}[] -\rail@nont{\isa{method}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{method}}[] -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} refines the goal by proof - method \isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}; facts for forward chaining are passed if so - indicated by \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} mode. - - \item \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} refines any remaining goals by - proof method \isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} and concludes the sub-proof by assumption. - If the goal had been \isa{{\isaliteral{22}{\isachardoublequote}}show{\isaliteral{22}{\isachardoublequote}}} (or \isa{{\isaliteral{22}{\isachardoublequote}}thus{\isaliteral{22}{\isachardoublequote}}}), some - pending sub-goal is solved as well by the rule resulting from the - result \emph{exported} into the enclosing goal context. Thus \isa{{\isaliteral{22}{\isachardoublequote}}qed{\isaliteral{22}{\isachardoublequote}}} may fail for two reasons: either \isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} fails, or the - resulting rule does not fit to any pending goal\footnote{This - includes any additional ``strong'' assumptions as introduced by - \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}.} of the enclosing context. Debugging such a - situation might involve temporarily changing \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} into - \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, or weakening the local context by replacing - occurrences of \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} by \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}. - - \item \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} is a \emph{terminal - proof}\index{proof!terminal}; it abbreviates \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}, but with - backtracking across both methods. Debugging an unsuccessful - \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} command can be done by expanding its - definition; in many cases \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} (or even - \isa{{\isaliteral{22}{\isachardoublequote}}apply{\isaliteral{22}{\isachardoublequote}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}) is already sufficient to see the - problem. - - \item ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' is a \emph{default - proof}\index{proof!default}; it abbreviates \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}rule{\isaliteral{22}{\isachardoublequote}}}. - - \item ``\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}'' is a \emph{trivial - proof}\index{proof!trivial}; it abbreviates \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}this{\isaliteral{22}{\isachardoublequote}}}. - - \item \hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}} is a \emph{fake proof}\index{proof!fake} - pretending to solve the pending claim without further ado. This - only works in interactive development, or if the \verb|quick_and_dirty| flag is enabled (in ML). Facts emerging from fake - proofs are not the real thing. Internally, each theorem container - is tainted by an oracle invocation, which is indicated as ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{21}{\isacharbang}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}'' in the printed result. - - The most important application of \hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}} is to support - experimentation and top-down proof development. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Fundamental methods and attributes \label{sec:pure-meth-att}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The following proof methods and attributes refer to basic logical - operations of Isar. Further methods and attributes are provided by - several generic and object-logic specific tools and packages (see - \chref{ch:gen-tools} and \chref{ch:hol}). - - \begin{matharray}{rcl} - \indexdef{}{method}{-}\hypertarget{method.-}{\hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}}} & : & \isa{method} \\ - \indexdef{}{method}{fact}\hypertarget{method.fact}{\hyperlink{method.fact}{\mbox{\isa{fact}}}} & : & \isa{method} \\ - \indexdef{}{method}{assumption}\hypertarget{method.assumption}{\hyperlink{method.assumption}{\mbox{\isa{assumption}}}} & : & \isa{method} \\ - \indexdef{}{method}{this}\hypertarget{method.this}{\hyperlink{method.this}{\mbox{\isa{this}}}} & : & \isa{method} \\ - \indexdef{Pure}{method}{rule}\hypertarget{method.Pure.rule}{\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}} & : & \isa{method} \\ - \indexdef{Pure}{attribute}{intro}\hypertarget{attribute.Pure.intro}{\hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}} & : & \isa{attribute} \\ - \indexdef{Pure}{attribute}{elim}\hypertarget{attribute.Pure.elim}{\hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}} & : & \isa{attribute} \\ - \indexdef{Pure}{attribute}{dest}\hypertarget{attribute.Pure.dest}{\hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}}} & : & \isa{attribute} \\ - \indexdef{Pure}{attribute}{rule}\hypertarget{attribute.Pure.rule}{\hyperlink{attribute.Pure.rule}{\mbox{\isa{rule}}}} & : & \isa{attribute} \\[0.5ex] - \indexdef{}{attribute}{OF}\hypertarget{attribute.OF}{\hyperlink{attribute.OF}{\mbox{\isa{OF}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{of}\hypertarget{attribute.of}{\hyperlink{attribute.of}{\mbox{\isa{of}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{where}\hypertarget{attribute.where}{\hyperlink{attribute.where}{\mbox{\isa{where}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{method.fact}{\mbox{\isa{fact}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@endbar -\rail@end -\rail@begin{4}{\isa{rulemod}} -\rail@bar -\rail@term{\isa{intro}}[] -\rail@nextbar{1} -\rail@term{\isa{elim}}[] -\rail@nextbar{2} -\rail@term{\isa{dest}}[] -\rail@endbar -\rail@bar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{1} -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@nextbar{3} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{3}{} -\rail@bar -\rail@term{\hyperlink{attribute.intro}{\mbox{\isa{intro}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{attribute.elim}{\mbox{\isa{elim}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{attribute.dest}{\mbox{\isa{dest}}}}[] -\rail@endbar -\rail@bar -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@nextbar{1} -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.Pure.rule}{\mbox{\isa{rule}}}}[] -\rail@term{\isa{del}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.OF}{\mbox{\isa{OF}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.of}{\mbox{\isa{of}}}}[] -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{concl}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@endbar -\rail@end -\rail@begin{6}{} -\rail@term{\hyperlink{attribute.where}{\mbox{\isa{where}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.var}{\mbox{\isa{var}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.typefree}{\mbox{\isa{typefree}}}}[] -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.typevar}{\mbox{\isa{typevar}}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item ``\hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}}'' (minus) does nothing but insert the forward - chaining facts as premises into the goal. Note that command - \indexref{}{command}{proof}\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} without any method actually performs a single - reduction step using the \indexref{Pure}{method}{rule}\hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method; thus a plain - \emph{do-nothing} proof step would be ``\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' rather than \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} alone. - - \item \hyperlink{method.fact}{\mbox{\isa{fact}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} composes some fact from - \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} (or implicitly from the current proof context) - modulo unification of schematic type and term variables. The rule - structure is not taken into account, i.e.\ meta-level implication is - considered atomic. This is the same principle underlying literal - facts (cf.\ \secref{sec:syn-att}): ``\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{fact}'' is equivalent to ``\hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\verb|`|\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}}\verb|`|'' provided that - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} is an instance of some known \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} in the - proof context. - - \item \hyperlink{method.assumption}{\mbox{\isa{assumption}}} solves some goal by a single assumption - step. All given facts are guaranteed to participate in the - refinement; this means there may be only 0 or 1 in the first place. - Recall that \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} (\secref{sec:proof-steps}) already - concludes any remaining sub-goals by assumption, so structured - proofs usually need not quote the \hyperlink{method.assumption}{\mbox{\isa{assumption}}} method at - all. - - \item \hyperlink{method.this}{\mbox{\isa{this}}} applies all of the current facts directly as - rules. Recall that ``\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}'' (dot) abbreviates ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{this}''. - - \item \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} applies some rule given as - argument in backward manner; facts are used to reduce the rule - before applying it to the goal. Thus \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} without facts - is plain introduction, while with facts it becomes elimination. - - When no arguments are given, the \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}} method tries to pick - appropriate rules automatically, as declared in the current context - using the \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, - \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}} attributes (see below). This is the - default behavior of \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} and ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' - (double-dot) steps (see \secref{sec:proof-steps}). - - \item \hyperlink{attribute.Pure.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.Pure.elim}{\mbox{\isa{elim}}}, and - \hyperlink{attribute.Pure.dest}{\mbox{\isa{dest}}} declare introduction, elimination, and - destruct rules, to be used with method \hyperlink{method.Pure.rule}{\mbox{\isa{rule}}}, and similar - tools. Note that the latter will ignore rules declared with - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'', while ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}}'' are used most aggressively. - - The classical reasoner (see \secref{sec:classical}) introduces its - own variants of these attributes; use qualified names to access the - present versions of Isabelle/Pure, i.e.\ \hyperlink{attribute.Pure.Pure.intro}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}intro}}}. - - \item \hyperlink{attribute.Pure.rule}{\mbox{\isa{rule}}}~\isa{del} undeclares introduction, - elimination, or destruct rules. - - \item \hyperlink{attribute.OF}{\mbox{\isa{OF}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} applies some theorem to all - of the given rules \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} in canonical right-to-left - order, which means that premises stemming from the \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} - emerge in parallel in the result, without interfering with each - other. In many practical situations, the \isa{{\isaliteral{22}{\isachardoublequote}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} do not have - premises themselves, so \isa{{\isaliteral{22}{\isachardoublequote}}rule\ {\isaliteral{5B}{\isacharbrackleft}}OF\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} can be actually - read as functional application (modulo unification). - - Argument positions may be effectively skipped by using ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' - (underscore), which refers to the propositional identity rule in the - Pure theory. - - \item \hyperlink{attribute.of}{\mbox{\isa{of}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} performs positional - instantiation of term variables. The terms \isa{{\isaliteral{22}{\isachardoublequote}}t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} are - substituted for any schematic variables occurring in a theorem from - left to right; ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' (underscore) indicates to skip a - position. Arguments following a ``\isa{{\isaliteral{22}{\isachardoublequote}}concl{\isaliteral{3A}{\isacharcolon}}{\isaliteral{22}{\isachardoublequote}}}'' specification - refer to positions of the conclusion of a rule. - - \item \hyperlink{attribute.where}{\mbox{\isa{where}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3D}{\isacharequal}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} - performs named instantiation of schematic type and term variables - occurring in a theorem. Schematic variables have to be specified on - the left-hand side (e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isadigit{1}}{\isaliteral{2E}{\isachardot}}{\isadigit{3}}{\isaliteral{22}{\isachardoublequote}}}). The question mark may - be omitted if the variable name is a plain identifier without index. - As type instantiations are inferred from term instantiations, - explicit type instantiations are seldom necessary. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Emulating tactic scripts \label{sec:tactic-commands}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Isar provides separate commands to accommodate tactic-style - proof scripts within the same system. While being outside the - orthodox Isar proof language, these might come in handy for - interactive exploration and debugging, or even actual tactical proof - within new-style theories (to benefit from document preparation, for - example). See also \secref{sec:tactics} for actual tactics, that - have been encapsulated as proof methods. Proper proof methods may - be used in scripts, too. - - \begin{matharray}{rcl} - \indexdef{}{command}{apply}\hypertarget{command.apply}{\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{apply\_end}\hypertarget{command.apply-end}{\hyperlink{command.apply-end}{\mbox{\isa{\isacommand{apply{\isaliteral{5F}{\isacharunderscore}}end}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{done}\hypertarget{command.done}{\hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{7C}{\isacharbar}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{defer}\hypertarget{command.defer}{\hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{prefer}\hypertarget{command.prefer}{\hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{back}\hypertarget{command.back}{\hyperlink{command.back}{\mbox{\isa{\isacommand{back}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.apply-end}{\mbox{\isa{\isacommand{apply{\isaliteral{5F}{\isacharunderscore}}end}}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.method}{\mbox{\isa{method}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{m} applies proof method \isa{m} in - initial position, but unlike \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}} it retains ``\isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' mode. Thus consecutive method applications may be - given just as in tactic scripts. - - Facts are passed to \isa{m} as indicated by the goal's - forward-chain mode, and are \emph{consumed} afterwards. Thus any - further \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}} command would always work in a purely - backward manner. - - \item \hyperlink{command.apply-end}{\mbox{\isa{\isacommand{apply{\isaliteral{5F}{\isacharunderscore}}end}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m{\isaliteral{22}{\isachardoublequote}}} applies proof method \isa{m} as if in terminal position. Basically, this simulates a - multi-step tactic script for \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}, but may be given - anywhere within the proof body. - - No facts are passed to \isa{m} here. Furthermore, the static - context is that of the enclosing goal (as for actual \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}). Thus the proof method may not refer to any assumptions - introduced in the current body, for example. - - \item \hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}} completes a proof script, provided that the - current goal state is solved completely. Note that actual - structured proof commands (e.g.\ ``\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}'' or \hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}}) may be used to conclude proof scripts as well. - - \item \hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}}~\isa{n} and \hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}}~\isa{n} - shuffle the list of pending goals: \hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}} puts off - sub-goal \isa{n} to the end of the list (\isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} by - default), while \hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}} brings sub-goal \isa{n} to the - front. - - \item \hyperlink{command.back}{\mbox{\isa{\isacommand{back}}}} does back-tracking over the result sequence - of the latest proof command. Basically, any proof command may - return multiple results. - - \end{description} - - Any proper Isar proof method may be used with tactic script commands - such as \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}. A few additional emulations of actual - tactics are provided as well; these would be never used in actual - structured proofs, of course.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Defining proof methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{method\_setup}\hypertarget{command.method-setup}{\hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.method-setup}{\mbox{\isa{\isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ text\ description{\isaliteral{22}{\isachardoublequote}}} - defines a proof method in the current theory. The given \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} has to be an ML expression of type - \verb|(Proof.context -> Proof.method) context_parser|, cf.\ - basic parsers defined in structure \verb|Args| and \verb|Attrib|. There are also combinators like \verb|METHOD| and \verb|SIMPLE_METHOD| to turn certain tactic forms into official proof - methods; the primed versions refer to tactics with explicit goal - addressing. - - Here are some example method definitions: - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -\ \ % -\endisadelimML -% -\isatagML -\isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ my{\isaliteral{5F}{\isacharunderscore}}method{\isadigit{1}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ \ \ Scan{\isaliteral{2E}{\isachardot}}succeed\ {\isaliteral{28}{\isacharparenleft}}K\ {\isaliteral{28}{\isacharparenleft}}SIMPLE{\isaliteral{5F}{\isacharunderscore}}METHOD{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}fn\ i{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ no{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ {\isaliteral{2A7D}{\isacharverbatimclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}my\ first\ method\ {\isaliteral{28}{\isacharparenleft}}without\ any\ arguments{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ my{\isaliteral{5F}{\isacharunderscore}}method{\isadigit{2}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ \ \ Scan{\isaliteral{2E}{\isachardot}}succeed\ {\isaliteral{28}{\isacharparenleft}}fn\ ctxt{\isaliteral{3A}{\isacharcolon}}\ Proof{\isaliteral{2E}{\isachardot}}context\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ SIMPLE{\isaliteral{5F}{\isacharunderscore}}METHOD{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}fn\ i{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ no{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ {\isaliteral{2A7D}{\isacharverbatimclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}my\ second\ method\ {\isaliteral{28}{\isacharparenleft}}with\ context{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{method{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ my{\isaliteral{5F}{\isacharunderscore}}method{\isadigit{3}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ \ \ Attrib{\isaliteral{2E}{\isachardot}}thms\ {\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}fn\ thms{\isaliteral{3A}{\isacharcolon}}\ thm\ list\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ ctxt{\isaliteral{3A}{\isacharcolon}}\ Proof{\isaliteral{2E}{\isachardot}}context\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ SIMPLE{\isaliteral{5F}{\isacharunderscore}}METHOD{\isaliteral{27}{\isacharprime}}\ {\isaliteral{28}{\isacharparenleft}}fn\ i{\isaliteral{3A}{\isacharcolon}}\ int\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ no{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ {\isaliteral{2A7D}{\isacharverbatimclose}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}my\ third\ method\ {\isaliteral{28}{\isacharparenleft}}with\ theorem\ arguments\ and\ context{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupsection{Generalized elimination \label{sec:obtain}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{obtain}\hypertarget{command.obtain}{\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{guess}\hypertarget{command.guess}{\hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Generalized elimination means that additional elements with certain - properties may be introduced in the current context, by virtue of a - locally proven ``soundness statement''. Technically speaking, the - \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} language element is like a declaration of - \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}} and \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} (see also see - \secref{sec:proof-context}), together with a soundness proof of its - additional claim. According to the nature of existential reasoning, - assumptions get eliminated from any result exported from the context - later, provided that the corresponding parameters do \emph{not} - occur in the conclusion. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{\isakeyword{where}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - The derived Isar command \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} is defined as follows - (where \isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}} shall refer to (optional) - facts indicated for forward chaining). - \begin{matharray}{l} - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}using\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}}~~\hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ \ {\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}proof{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} \\[1ex] - \quad \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}thesis{\isaliteral{2E}{\isachardot}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\hyperlink{method.succeed}{\mbox{\isa{succeed}}} \\ - \qquad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{thesis} \\ - \qquad \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}that\ {\isaliteral{5B}{\isacharbrackleft}}Pure{\isaliteral{2E}{\isachardot}}intro{\isaliteral{3F}{\isacharquery}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} \\ - \qquad \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{thesis} \\ - \quad\qquad \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{{\isaliteral{2D}{\isacharminus}}} \\ - \quad\qquad \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ \ {\isaliteral{5C3C6C616E676C653E}{\isasymlangle}}proof{\isaliteral{5C3C72616E676C653E}{\isasymrangle}}{\isaliteral{22}{\isachardoublequote}}} \\ - \quad \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}} \\ - \quad \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}\ a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Typically, the soundness proof is relatively straight-forward, often - just by canonical automated tools such as ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{simp}'' or ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{blast}''. Accordingly, the - ``\isa{that}'' reduction above is declared as simplification and - introduction rule. - - In a sense, \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} represents at the level of Isar - proofs what would be meta-logical existential quantifiers and - conjunctions. This concept has a broad range of useful - applications, ranging from plain elimination (or introduction) of - object-level existential and conjunctions, to elimination over - results of symbolic evaluation of recursive definitions, for - example. Also note that \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} without parameters acts - much like \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, where the result is treated as a - genuine assumption. - - An alternative name to be used instead of ``\isa{that}'' above may - be given in parentheses. - - \medskip The improper variant \hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}} is similar to - \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}, but derives the obtained statement from the - course of reasoning! The proof starts with a fixed goal \isa{thesis}. The subsequent proof may refine this to anything of the - form like \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}}, but must not introduce new subgoals. The - final goal state is then used as reduction rule for the obtain - scheme described above. Obtained parameters \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} are marked as internal by default, which prevents the - proof context from being polluted by ad-hoc variables. The variable - names and type constraints given as arguments for \hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}} - specify a prefix of obtained parameters explicitly in the text. - - It is important to note that the facts introduced by \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}} and \hyperlink{command.guess}{\mbox{\isa{\isacommand{guess}}}} may not be polymorphic: any - type-variables occurring here are fixed in the present context!% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Calculational reasoning \label{sec:calculation}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{also}\hypertarget{command.also}{\hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{finally}\hypertarget{command.finally}{\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{moreover}\hypertarget{command.moreover}{\hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{ultimately}\hypertarget{command.ultimately}{\hyperlink{command.ultimately}{\mbox{\isa{\isacommand{ultimately}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_trans\_rules}\hypertarget{command.print-trans-rules}{\hyperlink{command.print-trans-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}trans{\isaliteral{5F}{\isacharunderscore}}rules}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.trans}{\mbox{\isa{trans}}} & : & \isa{attribute} \\ - \hyperlink{attribute.sym}{\mbox{\isa{sym}}} & : & \isa{attribute} \\ - \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} & : & \isa{attribute} \\ - \end{matharray} - - Calculational proof is forward reasoning with implicit application - of transitivity rules (such those of \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3D}{\isacharequal}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C653E}{\isasymle}}{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3C}{\isacharless}}{\isaliteral{22}{\isachardoublequote}}}). Isabelle/Isar maintains an auxiliary fact register - \indexref{}{fact}{calculation}\hyperlink{fact.calculation}{\mbox{\isa{calculation}}} for accumulating results obtained by - transitivity composed with the current result. Command \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} updates \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} involving \hyperlink{fact.this}{\mbox{\isa{this}}}, while - \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} exhibits the final \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} by - forward chaining towards the next goal statement. Both commands - require valid current facts, i.e.\ may occur only after commands - that produce theorems such as \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}, or some finished proof of \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}, \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} etc. The \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}} and \hyperlink{command.ultimately}{\mbox{\isa{\isacommand{ultimately}}}} - commands are similar to \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} and \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}, - but only collect further results in \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} without - applying any rules yet. - - Also note that the implicit term abbreviation ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' has - its canonical application with calculational proofs. It refers to - the argument of the preceding statement. (The argument of a curried - infix expression happens to be its right-hand side.) - - Isabelle/Isar calculations are implicitly subject to block structure - in the sense that new threads of calculational reasoning are - commenced for any new block (as opened by a local goal, for - example). This means that, apart from being able to nest - calculations, there is no separate \emph{begin-calculation} command - required. - - \medskip The Isar calculation proof commands may be defined as - follows:\footnote{We suppress internal bookkeeping such as proper - handling of block-structure.} - - \begin{matharray}{rcl} - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{2B}{\isacharplus}}{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ trans\ {\isaliteral{5B}{\isacharbrackleft}}OF\ calculation\ this{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} \\[0.5ex] - \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} & \equiv & \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\[0.5ex] - \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}} & \equiv & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ calculation\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.ultimately}{\mbox{\isa{\isacommand{ultimately}}}} & \equiv & \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.trans}{\mbox{\isa{trans}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{add}}[] -\rail@nextbar{2} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} maintains the auxiliary - \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} register as follows. The first occurrence of - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} in some calculational thread initializes \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} by \hyperlink{fact.this}{\mbox{\isa{this}}}. Any subsequent \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} on - the same level of block-structure updates \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} by - some transitivity rule applied to \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} and \hyperlink{fact.this}{\mbox{\isa{this}}} (in that order). Transitivity rules are picked from the - current context, unless alternative rules are given as explicit - arguments. - - \item \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}a\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ a\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} maintaining \hyperlink{fact.calculation}{\mbox{\isa{calculation}}} in the same way as \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}, and concludes the - current calculational thread. The final result is exhibited as fact - for forward chaining towards the next goal. Basically, \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} just abbreviates \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\hyperlink{fact.calculation}{\mbox{\isa{calculation}}}. Typical idioms for concluding - calculational proofs are ``\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isaliteral{3F}{\isacharquery}}thesis}~\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}'' and ``\hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}}~\hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}}''. - - \item \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}} and \hyperlink{command.ultimately}{\mbox{\isa{\isacommand{ultimately}}}} are - analogous to \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} and \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}, but collect - results only, without applying rules. - - \item \hyperlink{command.print-trans-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}trans{\isaliteral{5F}{\isacharunderscore}}rules}}}} prints the list of transitivity - rules (for calculational commands \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}} and \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}}) and symmetry rules (for the \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} - operation and single step elimination patters) of the current - context. - - \item \hyperlink{attribute.trans}{\mbox{\isa{trans}}} declares theorems as transitivity rules. - - \item \hyperlink{attribute.sym}{\mbox{\isa{sym}}} declares symmetry rules, as well as - \hyperlink{attribute.Pure.elim}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}elim}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} rules. - - \item \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} resolves a theorem with some rule - declared as \hyperlink{attribute.sym}{\mbox{\isa{sym}}} in the current context. For example, - ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{3A}{\isacharcolon}}\ x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}}'' produces a - swapped fact derived from that assumption. - - In structured proof texts it is often more appropriate to use an - explicit single-step elimination proof, such as ``\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}y\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}''. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proof by cases and induction \label{sec:cases-induct}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Rule contexts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{case}\hypertarget{command.case}{\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_cases}\hypertarget{command.print-cases}{\hyperlink{command.print-cases}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}cases}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{case\_names}\hypertarget{attribute.case-names}{\hyperlink{attribute.case-names}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}names}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{case\_conclusion}\hypertarget{attribute.case-conclusion}{\hyperlink{attribute.case-conclusion}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}conclusion}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{params}\hypertarget{attribute.params}{\hyperlink{attribute.params}{\mbox{\isa{params}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{consumes}\hypertarget{attribute.consumes}{\hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}} & : & \isa{attribute} \\ - \end{matharray} - - The puristic way to build up Isar proof contexts is by explicit - language elements like \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}, \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, - \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}} (see \secref{sec:proof-context}). This is adequate - for plain natural deduction, but easily becomes unwieldy in concrete - verification tasks, which typically involve big induction rules with - several cases. - - The \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}} command provides a shorthand to refer to a - local context symbolically: certain proof methods provide an - environment of named ``cases'' of the form \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{3A}{\isacharcolon}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}; the effect of ``\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{c}'' is then equivalent to ``\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}''. Term bindings may be covered as well, notably - \hyperlink{variable.?case}{\mbox{\isa{{\isaliteral{3F}{\isacharquery}}case}}} for the main conclusion. - - By default, the ``terminology'' \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} of - a case value is marked as hidden, i.e.\ there is no way to refer to - such parameters in the subsequent proof text. After all, original - rule parameters stem from somewhere outside of the current proof - text. By using the explicit form ``\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ y\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' instead, the proof author is able to - chose local names that fit nicely into the current context. - - \medskip It is important to note that proper use of \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}} does not provide means to peek at the current goal state, - which is not directly observable in Isar! Nonetheless, goal - refinement commands do provide named cases \isa{{\isaliteral{22}{\isachardoublequote}}goal\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} - for each subgoal \isa{{\isaliteral{22}{\isachardoublequote}}i\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ n{\isaliteral{22}{\isachardoublequote}}} of the resulting goal state. - Using this extra feature requires great care, because some bits of - the internal tactical machinery intrude the proof text. In - particular, parameter names stemming from the left-over of automated - reasoning tools are usually quite unpredictable. - - Under normal circumstances, the text of cases emerge from standard - elimination or induction rules, which in turn are derived from - previous theory specifications in a canonical way (say from - \hyperlink{command.inductive}{\mbox{\isa{\isacommand{inductive}}}} definitions). - - \medskip Proper cases are only available if both the proof method - and the rules involved support this. By using appropriate - attributes, case names, conclusions, and parameters may be also - declared by hand. Thus variant versions of rules that have been - derived manually become ready to use in advanced case analysis - later. - - \begin{railoutput} -\rail@begin{4}{} -\rail@term{\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}}[] -\rail@bar -\rail@nont{\isa{caseref}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\isa{caseref}}[] -\rail@plus -\rail@bar -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@nextplus{3} -\rail@endplus -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{caseref}} -\rail@nont{\isa{nameref}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{attributes}}[] -\rail@endbar -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{attribute.case-names}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}names}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] -\rail@plus -\rail@bar -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endbar -\rail@nextplus{3} -\rail@endplus -\rail@term{\isa{{\isaliteral{5D}{\isacharbrackright}}}}[] -\rail@endbar -\rail@nextplus{4} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.case-conclusion}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}conclusion}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{attribute.params}{\mbox{\isa{params}}}}[] -\rail@plus -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@endplus -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} invokes a named local - context \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{3A}{\isacharcolon}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}, as provided by an - appropriate proof method (such as \indexref{}{method}{cases}\hyperlink{method.cases}{\mbox{\isa{cases}}} and - \indexref{}{method}{induct}\hyperlink{method.induct}{\mbox{\isa{induct}}}). The command ``\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}c\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' abbreviates ``\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}''. - - \item \hyperlink{command.print-cases}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}cases}}}} prints all local contexts of the - current state, using Isar proof language notation. - - \item \hyperlink{attribute.case-names}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}names}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}} declares names for - the local contexts of premises of a theorem; \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}} - refers to the \emph{prefix} of the list of premises. Each of the - cases \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} can be of the form \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{5B}{\isacharbrackleft}}h\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ h\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} where - the \isa{{\isaliteral{22}{\isachardoublequote}}h\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ h\isaliteral{5C3C5E697375623E}{}\isactrlisub n{\isaliteral{22}{\isachardoublequote}}} are the names of the hypotheses in case \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub i{\isaliteral{22}{\isachardoublequote}}} - from left to right. - - \item \hyperlink{attribute.case-conclusion}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}conclusion}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ d\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ d\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}} declares - names for the conclusions of a named premise \isa{c}; here \isa{{\isaliteral{22}{\isachardoublequote}}d\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ d\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{22}{\isachardoublequote}}} refers to the prefix of arguments of a logical formula - built by nesting a binary connective (e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}). - - Note that proof methods such as \hyperlink{method.induct}{\mbox{\isa{induct}}} and \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} already provide a default name for the conclusion as a - whole. The need to name subformulas only arises with cases that - split into several sub-cases, as in common co-induction rules. - - \item \hyperlink{attribute.params}{\mbox{\isa{params}}}~\isa{{\isaliteral{22}{\isachardoublequote}}p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ q\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ q\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} renames - the innermost parameters of premises \isa{{\isaliteral{22}{\isachardoublequote}}{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ n{\isaliteral{22}{\isachardoublequote}}} of some - theorem. An empty list of names may be given to skip positions, - leaving the present parameters unchanged. - - Note that the default usage of case rules does \emph{not} directly - expose parameters to the proof context. - - \item \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}~\isa{n} declares the number of ``major - premises'' of a rule, i.e.\ the number of facts to be consumed when - it is applied by an appropriate proof method. The default value of - \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}} is \isa{{\isaliteral{22}{\isachardoublequote}}n\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, which is appropriate for - the usual kind of cases and induction rules for inductive sets (cf.\ - \secref{sec:hol-inductive}). Rules without any \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}} declaration given are treated as if \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}~\isa{{\isadigit{0}}} had been specified. - - Note that explicit \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}} declarations are only - rarely needed; this is already taken care of automatically by the - higher-level \hyperlink{attribute.cases}{\mbox{\isa{cases}}}, \hyperlink{attribute.induct}{\mbox{\isa{induct}}}, and - \hyperlink{attribute.coinduct}{\mbox{\isa{coinduct}}} declarations. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Proof methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{method}{cases}\hypertarget{method.cases}{\hyperlink{method.cases}{\mbox{\isa{cases}}}} & : & \isa{method} \\ - \indexdef{}{method}{induct}\hypertarget{method.induct}{\hyperlink{method.induct}{\mbox{\isa{induct}}}} & : & \isa{method} \\ - \indexdef{}{method}{induction}\hypertarget{method.induction}{\hyperlink{method.induction}{\mbox{\isa{induction}}}} & : & \isa{method} \\ - \indexdef{}{method}{coinduct}\hypertarget{method.coinduct}{\hyperlink{method.coinduct}{\mbox{\isa{coinduct}}}} & : & \isa{method} \\ - \end{matharray} - - The \hyperlink{method.cases}{\mbox{\isa{cases}}}, \hyperlink{method.induct}{\mbox{\isa{induct}}}, \hyperlink{method.induction}{\mbox{\isa{induction}}}, - and \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} - methods provide a uniform interface to common proof techniques over - datatypes, inductive predicates (or sets), recursive functions etc. - The corresponding rules may be specified and instantiated in a - casual manner. Furthermore, these methods provide named local - contexts that may be invoked via the \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}} proof command - within the subsequent proof text. This accommodates compact proof - texts even when reasoning about large specifications. - - The \hyperlink{method.induct}{\mbox{\isa{induct}}} method also provides some additional - infrastructure in order to be applicable to structure statements - (either using explicit meta-level connectives, or including facts - and parameters separately). This avoids cumbersome encoding of - ``strengthened'' inductive statements within the object-logic. - - Method \hyperlink{method.induction}{\mbox{\isa{induction}}} differs from \hyperlink{method.induct}{\mbox{\isa{induct}}} only in - the names of the facts in the local context invoked by the \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}} - command. - - \begin{railoutput} -\rail@begin{6}{} -\rail@term{\hyperlink{method.cases}{\mbox{\isa{cases}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@plus -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@end -\rail@begin{6}{} -\rail@bar -\rail@term{\hyperlink{method.induct}{\mbox{\isa{induct}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{method.induction}{\mbox{\isa{induction}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{no{\isaliteral{5F}{\isacharunderscore}}simp}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\isa{definsts}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@cr{4} -\rail@bar -\rail@nextbar{5} -\rail@nont{\isa{arbitrary}}[] -\rail@endbar -\rail@bar -\rail@nextbar{5} -\rail@nont{\isa{taking}}[] -\rail@endbar -\rail@bar -\rail@nextbar{5} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{method.coinduct}{\mbox{\isa{coinduct}}}}[] -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@nont{\isa{taking}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{rule}}[] -\rail@endbar -\rail@end -\rail@begin{5}{\isa{rule}} -\rail@bar -\rail@bar -\rail@term{\isa{type}}[] -\rail@nextbar{1} -\rail@term{\isa{pred}}[] -\rail@nextbar{2} -\rail@term{\isa{set}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@nextbar{3} -\rail@term{\isa{rule}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.thmref}{\mbox{\isa{thmref}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{4}{\isa{definst}} -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5C3C65717569763E}{\isasymequiv}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.inst}{\mbox{\isa{inst}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{\isa{definsts}} -\rail@plus -\rail@nextplus{1} -\rail@cnont{\isa{definst}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{arbitrary}} -\rail@term{\isa{arbitrary}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@plus -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endplus -\rail@term{\isa{\isakeyword{and}}}[] -\rail@nextplus{2} -\rail@endplus -\rail@end -\rail@begin{1}{\isa{taking}} -\rail@term{\isa{taking}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.insts}{\mbox{\isa{insts}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{method.cases}{\mbox{\isa{cases}}}~\isa{{\isaliteral{22}{\isachardoublequote}}insts\ R{\isaliteral{22}{\isachardoublequote}}} applies method \hyperlink{method.rule}{\mbox{\isa{rule}}} with an appropriate case distinction theorem, instantiated to - the subjects \isa{insts}. Symbolic case names are bound according - to the rule's local contexts. - - The rule is determined as follows, according to the facts and - arguments passed to the \hyperlink{method.cases}{\mbox{\isa{cases}}} method: - - \medskip - \begin{tabular}{llll} - facts & & arguments & rule \\\hline - & \hyperlink{method.cases}{\mbox{\isa{cases}}} & & classical case split \\ - & \hyperlink{method.cases}{\mbox{\isa{cases}}} & \isa{t} & datatype exhaustion (type of \isa{t}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A\ t{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.cases}{\mbox{\isa{cases}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & inductive predicate/set elimination (of \isa{A}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.cases}{\mbox{\isa{cases}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ rule{\isaliteral{3A}{\isacharcolon}}\ R{\isaliteral{22}{\isachardoublequote}}} & explicit rule \isa{R} \\ - \end{tabular} - \medskip - - Several instantiations may be given, referring to the \emph{suffix} - of premises of the case rule; within each premise, the \emph{prefix} - of variables is instantiated. In most situations, only a single - term needs to be specified; this refers to the first variable of the - last premise (it is usually the same for all cases). The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}simp{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option can be used to disable pre-simplification of - cases (see the description of \hyperlink{method.induct}{\mbox{\isa{induct}}} below for details). - - \item \hyperlink{method.induct}{\mbox{\isa{induct}}}~\isa{{\isaliteral{22}{\isachardoublequote}}insts\ R{\isaliteral{22}{\isachardoublequote}}} and - \hyperlink{method.induction}{\mbox{\isa{induction}}}~\isa{{\isaliteral{22}{\isachardoublequote}}insts\ R{\isaliteral{22}{\isachardoublequote}}} are analogous to the - \hyperlink{method.cases}{\mbox{\isa{cases}}} method, but refer to induction rules, which are - determined as follows: - - \medskip - \begin{tabular}{llll} - facts & & arguments & rule \\\hline - & \hyperlink{method.induct}{\mbox{\isa{induct}}} & \isa{{\isaliteral{22}{\isachardoublequote}}P\ x{\isaliteral{22}{\isachardoublequote}}} & datatype induction (type of \isa{x}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7475726E7374696C653E}{\isasymturnstile}}\ A\ x{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.induct}{\mbox{\isa{induct}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & predicate/set induction (of \isa{A}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.induct}{\mbox{\isa{induct}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ rule{\isaliteral{3A}{\isacharcolon}}\ R{\isaliteral{22}{\isachardoublequote}}} & explicit rule \isa{R} \\ - \end{tabular} - \medskip - - Several instantiations may be given, each referring to some part of - a mutual inductive definition or datatype --- only related partial - induction rules may be used together, though. Any of the lists of - terms \isa{{\isaliteral{22}{\isachardoublequote}}P{\isaliteral{2C}{\isacharcomma}}\ x{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} refers to the \emph{suffix} of variables - present in the induction rule. This enables the writer to specify - only induction variables, or both predicates and variables, for - example. - - Instantiations may be definitional: equations \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} - introduce local definitions, which are inserted into the claim and - discharged after applying the induction rule. Equalities reappear - in the inductive cases, but have been transformed according to the - induction principle being involved here. In order to achieve - practically useful induction hypotheses, some variables occurring in - \isa{t} need to be fixed (see below). Instantiations of the form - \isa{t}, where \isa{t} is not a variable, are taken as a - shorthand for \mbox{\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}}, where \isa{x} is a fresh - variable. If this is not intended, \isa{t} has to be enclosed in - parentheses. By default, the equalities generated by definitional - instantiations are pre-simplified using a specific set of rules, - usually consisting of distinctness and injectivity theorems for - datatypes. This pre-simplification may cause some of the parameters - of an inductive case to disappear, or may even completely delete - some of the inductive cases, if one of the equalities occurring in - their premises can be simplified to \isa{False}. The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}no{\isaliteral{5F}{\isacharunderscore}}simp{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option can be used to disable pre-simplification. - Additional rules to be used in pre-simplification can be declared - using the \indexdef{}{attribute}{induct\_simp}\hypertarget{attribute.induct-simp}{\hyperlink{attribute.induct-simp}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}simp}}}} attribute. - - The optional ``\isa{{\isaliteral{22}{\isachardoublequote}}arbitrary{\isaliteral{3A}{\isacharcolon}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}}'' - specification generalizes variables \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} of the original goal before applying induction. One can - separate variables by ``\isa{{\isaliteral{22}{\isachardoublequote}}and{\isaliteral{22}{\isachardoublequote}}}'' to generalize them in other - goals then the first. Thus induction hypotheses may become - sufficiently general to get the proof through. Together with - definitional instantiations, one may effectively perform induction - over expressions of a certain structure. - - The optional ``\isa{{\isaliteral{22}{\isachardoublequote}}taking{\isaliteral{3A}{\isacharcolon}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}'' - specification provides additional instantiations of a prefix of - pending variables in the rule. Such schematic induction rules - rarely occur in practice, though. - - \item \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}}~\isa{{\isaliteral{22}{\isachardoublequote}}inst\ R{\isaliteral{22}{\isachardoublequote}}} is analogous to the - \hyperlink{method.induct}{\mbox{\isa{induct}}} method, but refers to coinduction rules, which are - determined as follows: - - \medskip - \begin{tabular}{llll} - goal & & arguments & rule \\\hline - & \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} & \isa{x} & type coinduction (type of \isa{x}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & predicate/set coinduction (of \isa{A}) \\ - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ rule{\isaliteral{3A}{\isacharcolon}}\ R{\isaliteral{22}{\isachardoublequote}}} & explicit rule \isa{R} \\ - \end{tabular} - - Coinduction is the dual of induction. Induction essentially - eliminates \isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}} towards a generic result \isa{{\isaliteral{22}{\isachardoublequote}}P\ x{\isaliteral{22}{\isachardoublequote}}}, - while coinduction introduces \isa{{\isaliteral{22}{\isachardoublequote}}A\ x{\isaliteral{22}{\isachardoublequote}}} starting with \isa{{\isaliteral{22}{\isachardoublequote}}B\ x{\isaliteral{22}{\isachardoublequote}}}, for a suitable ``bisimulation'' \isa{B}. The cases of a - coinduct rule are typically named after the predicates or sets being - covered, while the conclusions consist of several alternatives being - named after the individual destructor patterns. - - The given instantiation refers to the \emph{suffix} of variables - occurring in the rule's major premise, or conclusion if unavailable. - An additional ``\isa{{\isaliteral{22}{\isachardoublequote}}taking{\isaliteral{3A}{\isacharcolon}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ t\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}'' - specification may be required in order to specify the bisimulation - to be used in the coinduction step. - - \end{description} - - Above methods produce named local contexts, as determined by the - instantiated rule as given in the text. Beyond that, the \hyperlink{method.induct}{\mbox{\isa{induct}}} and \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} methods guess further instantiations - from the goal specification itself. Any persisting unresolved - schematic variables of the resulting rule will render the the - corresponding case invalid. The term binding \hyperlink{variable.?case}{\mbox{\isa{{\isaliteral{3F}{\isacharquery}}case}}} for - the conclusion will be provided with each case, provided that term - is fully specified. - - The \hyperlink{command.print-cases}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}cases}}}} command prints all named cases present - in the current proof state. - - \medskip Despite the additional infrastructure, both \hyperlink{method.cases}{\mbox{\isa{cases}}} - and \hyperlink{method.coinduct}{\mbox{\isa{coinduct}}} merely apply a certain rule, after - instantiation, while conforming due to the usual way of monotonic - natural deduction: the context of a structured statement \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} - reappears unchanged after the case split. - - The \hyperlink{method.induct}{\mbox{\isa{induct}}} method is fundamentally different in this - respect: the meta-level structure is passed through the - ``recursive'' course involved in the induction. Thus the original - statement is basically replaced by separate copies, corresponding to - the induction hypotheses and conclusion; the original goal context - is no longer available. Thus local assumptions, fixed parameters - and definitions effectively participate in the inductive rephrasing - of the original statement. - - In \hyperlink{method.induct}{\mbox{\isa{induct}}} proofs, local assumptions introduced by cases are split - into two different kinds: \isa{hyps} stemming from the rule and - \isa{prems} from the goal statement. This is reflected in the - extracted cases accordingly, so invoking ``\hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{c}'' will provide separate facts \isa{c{\isaliteral{2E}{\isachardot}}hyps} and \isa{c{\isaliteral{2E}{\isachardot}}prems}, - as well as fact \isa{c} to hold the all-inclusive list. - - In \hyperlink{method.induction}{\mbox{\isa{induction}}} proofs, local assumptions introduced by cases are - split into three different kinds: \isa{IH}, the induction hypotheses, - \isa{hyps}, the remaining hypotheses stemming from the rule, and - \isa{prems}, the assumptions from the goal statement. The names are - \isa{c{\isaliteral{2E}{\isachardot}}IH}, \isa{c{\isaliteral{2E}{\isachardot}}hyps} and \isa{c{\isaliteral{2E}{\isachardot}}prems}, as above. - - - \medskip Facts presented to either method are consumed according to - the number of ``major premises'' of the rule involved, which is - usually 0 for plain cases and induction rules of datatypes etc.\ and - 1 for rules of inductive predicates or sets and the like. The - remaining facts are inserted into the goal verbatim before the - actual \isa{cases}, \isa{induct}, or \isa{coinduct} rule is - applied.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Declaring rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{print\_induct\_rules}\hypertarget{command.print-induct-rules}{\hyperlink{command.print-induct-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}induct{\isaliteral{5F}{\isacharunderscore}}rules}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{cases}\hypertarget{attribute.cases}{\hyperlink{attribute.cases}{\mbox{\isa{cases}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{induct}\hypertarget{attribute.induct}{\hyperlink{attribute.induct}{\mbox{\isa{induct}}}} & : & \isa{attribute} \\ - \indexdef{}{attribute}{coinduct}\hypertarget{attribute.coinduct}{\hyperlink{attribute.coinduct}{\mbox{\isa{coinduct}}}} & : & \isa{attribute} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.cases}{\mbox{\isa{cases}}}}[] -\rail@nont{\isa{spec}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.induct}{\mbox{\isa{induct}}}}[] -\rail@nont{\isa{spec}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{attribute.coinduct}{\mbox{\isa{coinduct}}}}[] -\rail@nont{\isa{spec}}[] -\rail@end -\rail@begin{4}{\isa{spec}} -\rail@bar -\rail@bar -\rail@term{\isa{type}}[] -\rail@nextbar{1} -\rail@term{\isa{pred}}[] -\rail@nextbar{2} -\rail@term{\isa{set}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{3} -\rail@term{\isa{del}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.print-induct-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}induct{\isaliteral{5F}{\isacharunderscore}}rules}}}} prints cases and induct rules - for predicates (or sets) and types of the current context. - - \item \hyperlink{attribute.cases}{\mbox{\isa{cases}}}, \hyperlink{attribute.induct}{\mbox{\isa{induct}}}, and \hyperlink{attribute.coinduct}{\mbox{\isa{coinduct}}} (as attributes) declare rules for reasoning about - (co)inductive predicates (or sets) and types, using the - corresponding methods of the same name. Certain definitional - packages of object-logics usually declare emerging cases and - induction rules as expected, so users rarely need to intervene. - - Rules may be deleted via the \isa{{\isaliteral{22}{\isachardoublequote}}del{\isaliteral{22}{\isachardoublequote}}} specification, which - covers all of the \isa{{\isaliteral{22}{\isachardoublequote}}type{\isaliteral{22}{\isachardoublequote}}}/\isa{{\isaliteral{22}{\isachardoublequote}}pred{\isaliteral{22}{\isachardoublequote}}}/\isa{{\isaliteral{22}{\isachardoublequote}}set{\isaliteral{22}{\isachardoublequote}}} - sub-categories simultaneously. For example, \hyperlink{attribute.cases}{\mbox{\isa{cases}}}~\isa{del} removes any \hyperlink{attribute.cases}{\mbox{\isa{cases}}} rules declared for - some type, predicate, or set. - - Manual rule declarations usually refer to the \hyperlink{attribute.case-names}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}names}}} and \hyperlink{attribute.params}{\mbox{\isa{params}}} attributes to adjust names of - cases and parameters of a rule; the \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}} - declaration is taken care of automatically: \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}~\isa{{\isadigit{0}}} is specified for ``type'' rules and \hyperlink{attribute.consumes}{\mbox{\isa{consumes}}}~\isa{{\isadigit{1}}} for ``predicate'' / ``set'' rules. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Quick_Reference.tex --- a/doc-src/IsarRef/Thy/document/Quick_Reference.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,290 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Quick{\isaliteral{5F}{\isacharunderscore}}Reference}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Quick{\isaliteral{5F}{\isacharunderscore}}Reference\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Isabelle/Isar quick reference \label{ap:refcard}% -} -\isamarkuptrue% -% -\isamarkupsection{Proof commands% -} -\isamarkuptrue% -% -\isamarkupsubsection{Primitives and basic syntax% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x} & augment context by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C626F783E}{\isasymbox}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} & augment context by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7068693E}{\isasymphi}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C626F783E}{\isasymbox}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} & indicate forward chaining of facts \\ - \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} & prove local result \\ - \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} & prove local result, refining some goal \\ - \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{a} & indicate use of additional facts \\ - \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{a} & unfold definitional equations \\ - \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\dots~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} & indicate proof structure and refinements \\ - \hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}} & indicate explicit blocks \\ - \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} & switch blocks \\ - \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequote}}} & reconsider facts \\ - \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}p\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} & abbreviate terms by higher-order matching \\ - \hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ \ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} & declare local mixfix syntax \\ - \end{tabular} - - \medskip - - \begin{tabular}{rcl} - \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{22}{\isachardoublequote}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}prfx\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}\ stmt\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}method\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}prfx\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}} \\ - \isa{prfx} & = & \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{method} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}facts{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.unfolding}{\mbox{\isa{\isacommand{unfolding}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}facts{\isaliteral{22}{\isachardoublequote}}} \\ - \isa{stmt} & = & \hyperlink{command.braceleft}{\mbox{\isa{\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}stmt\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.braceright}{\mbox{\isa{\isacommand{{\isaliteral{7D}{\isacharbraceright}}}}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ facts{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.let}{\mbox{\isa{\isacommand{let}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}term\ {\isaliteral{3D}{\isacharequal}}\ term{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.write}{\mbox{\isa{\isacommand{write}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{28}{\isacharparenleft}}mixfix{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}var\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2B}{\isacharplus}}{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}~\isa{goal} \\ - \isa{goal} & = & \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ props\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Abbreviations and synonyms% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{rcl} - \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}m\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{rule} \\ - \hyperlink{command.dot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{this} \\ - \hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} \\ - \hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} \\ - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{a} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{a}~\hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} \\ - \hyperlink{command.with}{\mbox{\isa{\isacommand{with}}}}~\isa{a} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.then}{\mbox{\isa{\isacommand{then}}}} \\ - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this}~\hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.hence}{\mbox{\isa{\isacommand{hence}}}} \\ - \hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{this}~\hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C65717569763E}{\isasymequiv}}{\isaliteral{22}{\isachardoublequote}}} & \hyperlink{command.thus}{\mbox{\isa{\isacommand{thus}}}} \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Derived elements% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{rcl} - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isaliteral{2B}{\isacharplus}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ trans\ {\isaliteral{5B}{\isacharbrackleft}}OF\ calculation\ this{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.finally}{\mbox{\isa{\isacommand{finally}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.also}{\mbox{\isa{\isacommand{also}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\[0.5ex] - \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}calculation\ {\isaliteral{3D}{\isacharequal}}\ calculation\ this{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.ultimately}{\mbox{\isa{\isacommand{ultimately}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.moreover}{\mbox{\isa{\isacommand{moreover}}}}~\hyperlink{command.from}{\mbox{\isa{\isacommand{from}}}}~\isa{calculation} \\[0.5ex] - \hyperlink{command.presume}{\mbox{\isa{\isacommand{presume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}~\hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.case}{\mbox{\isa{\isacommand{case}}}}~\isa{c} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~\isa{x}~\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.sorry}{\mbox{\isa{\isacommand{sorry}}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C617070726F783E}{\isasymapprox}}{\isaliteral{22}{\isachardoublequote}}} & - \hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\isa{cheating} \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Diagnostic commands% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \hyperlink{command.pr}{\mbox{\isa{\isacommand{pr}}}} & print current state \\ - \hyperlink{command.thm}{\mbox{\isa{\isacommand{thm}}}}~\isa{a} & print fact \\ - \hyperlink{command.prop}{\mbox{\isa{\isacommand{prop}}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} & print proposition \\ - \hyperlink{command.term}{\mbox{\isa{\isacommand{term}}}}~\isa{t} & print term \\ - \hyperlink{command.typ}{\mbox{\isa{\isacommand{typ}}}}~\isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} & print type \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Proof methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Single steps (forward-chaining facts)}} \\[0.5ex] - \hyperlink{method.assumption}{\mbox{\isa{assumption}}} & apply some assumption \\ - \hyperlink{method.this}{\mbox{\isa{this}}} & apply current facts \\ - \hyperlink{method.rule}{\mbox{\isa{rule}}}~\isa{a} & apply some rule \\ - \hyperlink{method.rule}{\mbox{\isa{rule}}} & apply standard rule (default for \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}) \\ - \hyperlink{method.contradiction}{\mbox{\isa{contradiction}}} & apply \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}{\isaliteral{22}{\isachardoublequote}}} elimination rule (any order) \\ - \hyperlink{method.cases}{\mbox{\isa{cases}}}~\isa{t} & case analysis (provides cases) \\ - \hyperlink{method.induct}{\mbox{\isa{induct}}}~\isa{x} & proof by induction (provides cases) \\[2ex] - - \multicolumn{2}{l}{\textbf{Repeated steps (inserting facts)}} \\[0.5ex] - \hyperlink{method.-}{\mbox{\isa{{\isaliteral{2D}{\isacharminus}}}}} & no rules \\ - \hyperlink{method.intro}{\mbox{\isa{intro}}}~\isa{a} & introduction rules \\ - \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}} & class introduction rules \\ - \hyperlink{method.elim}{\mbox{\isa{elim}}}~\isa{a} & elimination rules \\ - \hyperlink{method.unfold}{\mbox{\isa{unfold}}}~\isa{a} & definitional rewrite rules \\[2ex] - - \multicolumn{2}{l}{\textbf{Automated proof tools (inserting facts)}} \\[0.5ex] - \hyperlink{method.iprover}{\mbox{\isa{iprover}}} & intuitionistic proof search \\ - \hyperlink{method.blast}{\mbox{\isa{blast}}}, \hyperlink{method.fast}{\mbox{\isa{fast}}} & Classical Reasoner \\ - \hyperlink{method.simp}{\mbox{\isa{simp}}}, \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} & Simplifier (+ Splitter) \\ - \hyperlink{method.auto}{\mbox{\isa{auto}}}, \hyperlink{method.force}{\mbox{\isa{force}}} & Simplifier + Classical Reasoner \\ - \hyperlink{method.arith}{\mbox{\isa{arith}}} & Arithmetic procedures \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Attributes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \multicolumn{2}{l}{\textbf{Rules}} \\[0.5ex] - \hyperlink{attribute.OF}{\mbox{\isa{OF}}}~\isa{a} & rule resolved with facts (skipping ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'') \\ - \hyperlink{attribute.of}{\mbox{\isa{of}}}~\isa{t} & rule instantiated with terms (skipping ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'') \\ - \hyperlink{attribute.where}{\mbox{\isa{where}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} & rule instantiated with terms, by variable name \\ - \hyperlink{attribute.symmetric}{\mbox{\isa{symmetric}}} & resolution with symmetry rule \\ - \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}~\isa{b} & resolution with another rule \\ - \hyperlink{attribute.rule-format}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}format}}} & result put into standard rule format \\ - \hyperlink{attribute.elim-format}{\mbox{\isa{elim{\isaliteral{5F}{\isacharunderscore}}format}}} & destruct rule turned into elimination rule format \\[1ex] - - \multicolumn{2}{l}{\textbf{Declarations}} \\[0.5ex] - \hyperlink{attribute.simp}{\mbox{\isa{simp}}} & Simplifier rule \\ - \hyperlink{attribute.intro}{\mbox{\isa{intro}}}, \hyperlink{attribute.elim}{\mbox{\isa{elim}}}, \hyperlink{attribute.dest}{\mbox{\isa{dest}}} & Pure or Classical Reasoner rule \\ - \hyperlink{attribute.iff}{\mbox{\isa{iff}}} & Simplifier + Classical Reasoner rule \\ - \hyperlink{attribute.split}{\mbox{\isa{split}}} & case split rule \\ - \hyperlink{attribute.trans}{\mbox{\isa{trans}}} & transitivity rule \\ - \hyperlink{attribute.sym}{\mbox{\isa{sym}}} & symmetry rule \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Rule declarations and methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{l|lllll} - & \hyperlink{method.rule}{\mbox{\isa{rule}}} & \hyperlink{method.iprover}{\mbox{\isa{iprover}}} & \hyperlink{method.blast}{\mbox{\isa{blast}}} & \hyperlink{method.simp}{\mbox{\isa{simp}}} & \hyperlink{method.auto}{\mbox{\isa{auto}}} \\ - & & & \hyperlink{method.fast}{\mbox{\isa{fast}}} & \hyperlink{method.simp-all}{\mbox{\isa{simp{\isaliteral{5F}{\isacharunderscore}}all}}} & \hyperlink{method.force}{\mbox{\isa{force}}} \\ - \hline - \hyperlink{attribute.Pure.elim}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}elim}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} \hyperlink{attribute.Pure.intro}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}intro}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.Pure.elim}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}elim}}} \hyperlink{attribute.Pure.intro}{\mbox{\isa{Pure{\isaliteral{2E}{\isachardot}}intro}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.elim}{\mbox{\isa{elim}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} \hyperlink{attribute.intro}{\mbox{\isa{intro}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.elim}{\mbox{\isa{elim}}} \hyperlink{attribute.intro}{\mbox{\isa{intro}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.iff}{\mbox{\isa{iff}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.iff}{\mbox{\isa{iff}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.elim}{\mbox{\isa{elim}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} \hyperlink{attribute.intro}{\mbox{\isa{intro}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}} - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.simp}{\mbox{\isa{simp}}} - & & & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.cong}{\mbox{\isa{cong}}} - & & & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{attribute.split}{\mbox{\isa{split}}} - & & & & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C74696D65733E}{\isasymtimes}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Emulating tactic scripts% -} -\isamarkuptrue% -% -\isamarkupsubsection{Commands% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \hyperlink{command.apply}{\mbox{\isa{\isacommand{apply}}}}~\isa{m} & apply proof method at initial position \\ - \hyperlink{command.apply-end}{\mbox{\isa{\isacommand{apply{\isaliteral{5F}{\isacharunderscore}}end}}}}~\isa{m} & apply proof method near terminal position \\ - \hyperlink{command.done}{\mbox{\isa{\isacommand{done}}}} & complete proof \\ - \hyperlink{command.defer}{\mbox{\isa{\isacommand{defer}}}}~\isa{n} & move subgoal to end \\ - \hyperlink{command.prefer}{\mbox{\isa{\isacommand{prefer}}}}~\isa{n} & move subgoal to beginning \\ - \hyperlink{command.back}{\mbox{\isa{\isacommand{back}}}} & backtrack last command \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Methods% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{tabular}{ll} - \hyperlink{method.rule-tac}{\mbox{\isa{rule{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{insts} & resolution (with instantiation) \\ - \hyperlink{method.erule-tac}{\mbox{\isa{erule{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{insts} & elim-resolution (with instantiation) \\ - \hyperlink{method.drule-tac}{\mbox{\isa{drule{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{insts} & destruct-resolution (with instantiation) \\ - \hyperlink{method.frule-tac}{\mbox{\isa{frule{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{insts} & forward-resolution (with instantiation) \\ - \hyperlink{method.cut-tac}{\mbox{\isa{cut{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{insts} & insert facts (with instantiation) \\ - \hyperlink{method.thin-tac}{\mbox{\isa{thin{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} & delete assumptions \\ - \hyperlink{method.subgoal-tac}{\mbox{\isa{subgoal{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} & new claims \\ - \hyperlink{method.rename-tac}{\mbox{\isa{rename{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{x} & rename innermost goal parameters \\ - \hyperlink{method.rotate-tac}{\mbox{\isa{rotate{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{n} & rotate assumptions of goal \\ - \hyperlink{method.tactic}{\mbox{\isa{tactic}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} & arbitrary ML tactic \\ - \hyperlink{method.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{t} & exhaustion (datatypes) \\ - \hyperlink{method.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{x} & induction (datatypes) \\ - \hyperlink{method.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}~\isa{t} & exhaustion + simplification (inductive predicates) \\ - \end{tabular}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Spec.tex --- a/doc-src/IsarRef/Thy/document/Spec.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2019 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Spec}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Spec\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Specifications% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Isabelle/Isar theory format integrates specifications and - proofs, supporting interactive development with unlimited undo - operation. There is an integrated document preparation system (see - \chref{ch:document-prep}), for typesetting formal developments - together with informal text. The resulting hyper-linked PDF - documents can be used both for WWW presentation and printed copies. - - The Isar proof language (see \chref{ch:proofs}) is embedded into the - theory language as a proper sub-language. Proof mode is entered by - stating some \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}} or \hyperlink{command.lemma}{\mbox{\isa{\isacommand{lemma}}}} at the theory - level, and left again with the final conclusion (e.g.\ via \hyperlink{command.qed}{\mbox{\isa{\isacommand{qed}}}}). Some theory specification mechanisms also require a proof, - such as \hyperlink{command.typedef}{\mbox{\isa{\isacommand{typedef}}}} in HOL, which demands non-emptiness of - the representing sets.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Defining theories \label{sec:begin-thy}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{theory}\hypertarget{command.theory}{\hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}toplevel\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{global}{command}{end}\hypertarget{command.global.end}{\hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ toplevel{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Isabelle/Isar theories are defined via theory files, which may - contain both specifications and proofs; occasionally definitional - mechanisms also require some explicit proof. The theory body may be - sub-structured by means of \emph{local theory targets}, such as - \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} and \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}. - - The first proper command of a theory is \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}, which - indicates imports of previous theories and optional dependencies on - other source files (usually in ML). Just preceding the initial - \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}} command there may be an optional \hyperlink{command.header}{\mbox{\isa{\isacommand{header}}}} declaration, which is only relevant to document - preparation: see also the other section markup commands in - \secref{sec:markup}. - - A theory is concluded by a final \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} command, - one that does not belong to a local theory target. No further - commands may follow such a global \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}}, - although some user-interfaces might pretend that trailing input is - admissible. - - \begin{railoutput} -\rail@begin{4}{} -\rail@term{\hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\isa{imports}}[] -\rail@cr{2} -\rail@bar -\rail@nextbar{3} -\rail@nont{\isa{keywords}}[] -\rail@endbar -\rail@bar -\rail@nextbar{3} -\rail@nont{\isa{uses}}[] -\rail@endbar -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{2}{\isa{imports}} -\rail@term{\isa{\isakeyword{imports}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{3}{\isa{keywords}} -\rail@term{\isa{\isakeyword{keywords}}}[] -\rail@plus -\rail@plus -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nont{\hyperlink{syntax.tags}{\mbox{\isa{tags}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{uses}} -\rail@term{\isa{\isakeyword{uses}}}[] -\rail@plus -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.parname}{\mbox{\isa{parname}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.theory}{\mbox{\isa{\isacommand{theory}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C494D504F5254533E}{\isasymIMPORTS}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} - starts a new theory \isa{A} based on the merge of existing - theories \isa{{\isaliteral{22}{\isachardoublequote}}B\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ B\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Due to the possibility to import more - than one ancestor, the resulting theory structure of an Isabelle - session forms a directed acyclic graph (DAG). Isabelle takes care - that sources contributing to the development graph are always - up-to-date: changed files are automatically rechecked whenever a - theory header specification is processed. - - The optional \indexdef{}{keyword}{keywords}\hypertarget{keyword.keywords}{\hyperlink{keyword.keywords}{\mbox{\isa{\isakeyword{keywords}}}}} specification declares outer - syntax (\chref{ch:outer-syntax}) that is introduced in this theory - later on (rare in end-user applications). Both minor keywords and - major keywords of the Isar command language need to be specified, in - order to make parsing of proof documents work properly. Command - keywords need to be classified according to their structural role in - the formal text. Examples may be seen in Isabelle/HOL sources - itself, such as \hyperlink{keyword.keywords}{\mbox{\isa{\isakeyword{keywords}}}}~\verb|"typedef"| - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ thy{\isaliteral{5F}{\isacharunderscore}}goal{\isaliteral{22}{\isachardoublequote}}} or \hyperlink{keyword.keywords}{\mbox{\isa{\isakeyword{keywords}}}}~\verb|"datatype"| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ thy{\isaliteral{5F}{\isacharunderscore}}decl{\isaliteral{22}{\isachardoublequote}}} for theory-level declarations - with and without proof, respectively. Additional \hyperlink{syntax.tags}{\mbox{\isa{tags}}} - provide defaults for document preparation (\secref{sec:tags}). - - The optional \indexdef{}{keyword}{uses}\hypertarget{keyword.uses}{\hyperlink{keyword.uses}{\mbox{\isa{\isakeyword{uses}}}}} specification declares additional - dependencies on external files (notably ML sources). Files will be - loaded immediately (as ML), unless the name is parenthesized. The - latter case records a dependency that needs to be resolved later in - the text, usually via explicit \indexref{}{command}{use}\hyperlink{command.use}{\mbox{\isa{\isacommand{use}}}} for ML files; - other file formats require specific load commands defined by the - corresponding tools or packages. - - \item \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} concludes the current theory - definition. Note that some other commands, e.g.\ local theory - targets \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} or \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}} may involve a - \hyperlink{keyword.begin}{\mbox{\isa{\isakeyword{begin}}}} that needs to be matched by \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}, according to the usual rules for nested blocks. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Local theory targets \label{sec:target}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{context}\hypertarget{command.context}{\hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{local}{command}{end}\hypertarget{command.local.end}{\hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - A local theory target is a context managed separately within the - enclosing theory. Contexts may introduce parameters (fixed - variables) and assumptions (hypotheses). Definitions and theorems - depending on the context may be added incrementally later on. - - \emph{Named contexts} refer to locales (cf.\ \secref{sec:locale}) or - type classes (cf.\ \secref{sec:class}); the name ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{22}{\isachardoublequote}}}'' - signifies the global theory context. - - \emph{Unnamed contexts} may introduce additional parameters and - assumptions, and results produced in the context are generalized - accordingly. Such auxiliary contexts may be nested within other - targets, like \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}, \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}, \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}, \hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}. - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\indexref{}{syntax}{includes}\hyperlink{syntax.includes}{\mbox{\isa{includes}}}}[] -\rail@endbar -\rail@plus -\rail@nextplus{1} -\rail@cnont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@endplus -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{1}{\indexdef{}{syntax}{target}\hypertarget{syntax.target}{\hyperlink{syntax.target}{\mbox{\isa{target}}}}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{in}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} opens a named - context, by recommencing an existing locale or class \isa{c}. - Note that locale and class definitions allow to include the - \hyperlink{keyword.begin}{\mbox{\isa{\isakeyword{begin}}}} keyword as well, in order to continue the local - theory immediately after the initial specification. - - \item \hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}bundles\ elements\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} opens - an unnamed context, by extending the enclosing global or local - theory target by the given declaration bundles (\secref{sec:bundle}) - and context elements (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C46495845533E}{\isasymFIXES}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C415353554D45533E}{\isasymASSUMES}}{\isaliteral{22}{\isachardoublequote}}} - etc.). This means any results stemming from definitions and proofs - in the extended context will be exported into the enclosing target - by lifting over extra parameters and premises. - - \item \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}} concludes the current local theory, - according to the nesting of contexts. Note that a global \hyperlink{command.global.end}{\mbox{\isa{\isacommand{end}}}} has a different meaning: it concludes the theory - itself (\secref{sec:begin-thy}). - - \item \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequote}}}\indexdef{}{keyword}{in}\hypertarget{keyword.in}{\hyperlink{keyword.in}{\mbox{\isa{\isakeyword{in}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} given after any - local theory command specifies an immediate target, e.g.\ - ``\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}'' or ``\hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ c{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}''. This works both in a local or - global theory context; the current target context will be suspended - for this command only. Note that ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C494E3E}{\isasymIN}}\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' will - always produce a global result independently of the current target - context. - - \end{description} - - The exact meaning of results produced within a local theory context - depends on the underlying target infrastructure (locale, type class - etc.). The general idea is as follows, considering a context named - \isa{c} with parameter \isa{x} and assumption \isa{{\isaliteral{22}{\isachardoublequote}}A{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}. - - Definitions are exported by introducing a global version with - additional arguments; a syntactic abbreviation links the long form - with the abstract version of the target context. For example, - \isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} becomes \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2E}{\isachardot}}a\ {\isaliteral{3F}{\isacharquery}}x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} at the theory - level (for arbitrary \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}}), together with a local - abbreviation \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c{\isaliteral{2E}{\isachardot}}a\ x{\isaliteral{22}{\isachardoublequote}}} in the target context (for the - fixed parameter \isa{x}). - - Theorems are exported by discharging the assumptions and - generalizing the parameters of the context. For example, \isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ B{\isaliteral{5B}{\isacharbrackleft}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} becomes \isa{{\isaliteral{22}{\isachardoublequote}}c{\isaliteral{2E}{\isachardot}}a{\isaliteral{3A}{\isacharcolon}}\ A{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}, again for arbitrary - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}x{\isaliteral{22}{\isachardoublequote}}}. - - \medskip The Isabelle/HOL library contains numerous applications of - locales and classes, e.g.\ see \verb|~~/src/HOL/Algebra|. An - example for an unnamed auxiliary contexts is given in \verb|~~/src/HOL/Isar_Examples/Group_Context.thy|.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Bundled declarations \label{sec:bundle}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{bundle}\hypertarget{command.bundle}{\hyperlink{command.bundle}{\mbox{\isa{\isacommand{bundle}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_bundles}\hypertarget{command.print-bundles}{\hyperlink{command.print-bundles}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}bundles}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{include}\hypertarget{command.include}{\hyperlink{command.include}{\mbox{\isa{\isacommand{include}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{including}\hypertarget{command.including}{\hyperlink{command.including}{\mbox{\isa{\isacommand{including}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{keyword}{includes}\hypertarget{keyword.includes}{\hyperlink{keyword.includes}{\mbox{\isa{\isakeyword{includes}}}}} & : & syntax \\ - \end{matharray} - - The outer syntax of fact expressions (\secref{sec:syn-att}) involves - theorems and attributes, which are evaluated in the context and - applied to it. Attributes may declare theorems to the context, as - in \isa{{\isaliteral{22}{\isachardoublequote}}this{\isaliteral{5F}{\isacharunderscore}}rule\ {\isaliteral{5B}{\isacharbrackleft}}intro{\isaliteral{5D}{\isacharbrackright}}\ that{\isaliteral{5F}{\isacharunderscore}}rule\ {\isaliteral{5B}{\isacharbrackleft}}elim{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} for example. - Configuration options (\secref{sec:config}) are special declaration - attributes that operate on the context without a theorem, as in - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}show{\isaliteral{5F}{\isacharunderscore}}types\ {\isaliteral{3D}{\isacharequal}}\ false{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} for example. - - Expressions of this form may be defined as \emph{bundled - declarations} in the context, and included in other situations later - on. Including declaration bundles augments a local context casually - without logical dependencies, which is in contrast to locales and - locale interpretation (\secref{sec:locale}). - - \begin{railoutput} -\rail@begin{6}{} -\rail@term{\hyperlink{command.bundle}{\mbox{\isa{\isacommand{bundle}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@cr{3} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@bar -\rail@term{\hyperlink{command.include}{\mbox{\isa{\isacommand{include}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.including}{\mbox{\isa{\isacommand{including}}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{\indexdef{}{syntax}{includes}\hypertarget{syntax.includes}{\hyperlink{syntax.includes}{\mbox{\isa{includes}}}}} -\rail@term{\isa{\isakeyword{includes}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.bundle}{\mbox{\isa{\isacommand{bundle}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\ {\isaliteral{3D}{\isacharequal}}\ decls{\isaliteral{22}{\isachardoublequote}}} defines a bundle of - declarations in the current context. The RHS is similar to the one - of the \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} command. Bundles defined in local theory - targets are subject to transformations via morphisms, when moved - into different application contexts; this works analogously to any - other local theory specification. - - \item \hyperlink{command.print-bundles}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}bundles}}}} prints the named bundles that are - available in the current context. - - \item \hyperlink{command.include}{\mbox{\isa{\isacommand{include}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} includes the declarations - from the given bundles into the current proof body context. This is - analogous to \hyperlink{command.note}{\mbox{\isa{\isacommand{note}}}} (\secref{sec:proof-facts}) with the - expanded bundles. - - \item \hyperlink{command.including}{\mbox{\isa{\isacommand{including}}}} is similar to \hyperlink{command.include}{\mbox{\isa{\isacommand{include}}}}, but - works in proof refinement (backward mode). This is analogous to - \hyperlink{command.using}{\mbox{\isa{\isacommand{using}}}} (\secref{sec:proof-facts}) with the expanded - bundles. - - \item \hyperlink{keyword.includes}{\mbox{\isa{\isakeyword{includes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} is similar to - \hyperlink{command.include}{\mbox{\isa{\isacommand{include}}}}, but works in situations where a specification - context is constructed, notably for \hyperlink{command.context}{\mbox{\isa{\isacommand{context}}}} and long - statements of \hyperlink{command.theorem}{\mbox{\isa{\isacommand{theorem}}}} etc. - - \end{description} - - Here is an artificial example of bundling various configuration - options:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{bundle}\isamarkupfalse% -\ trace\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5B}{\isacharbrackleft}}simp{\isaliteral{5F}{\isacharunderscore}}trace{\isaliteral{2C}{\isacharcomma}}\ blast{\isaliteral{5F}{\isacharunderscore}}trace{\isaliteral{2C}{\isacharcomma}}\ linarith{\isaliteral{5F}{\isacharunderscore}}trace{\isaliteral{2C}{\isacharcomma}}\ metis{\isaliteral{5F}{\isacharunderscore}}trace{\isaliteral{2C}{\isacharcomma}}\ smt{\isaliteral{5F}{\isacharunderscore}}trace{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{including}\isamarkupfalse% -\ trace% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ metis% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsection{Basic specification elements% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{axiomatization}\hypertarget{command.axiomatization}{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\ - \indexdef{}{command}{definition}\hypertarget{command.definition}{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{attribute}{defn}\hypertarget{attribute.defn}{\hyperlink{attribute.defn}{\mbox{\isa{defn}}}} & : & \isa{attribute} \\ - \indexdef{}{command}{abbreviation}\hypertarget{command.abbreviation}{\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_abbrevs}\hypertarget{command.print-abbrevs}{\hyperlink{command.print-abbrevs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ {\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - These specification mechanisms provide a slightly more abstract view - than the underlying primitives of \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}, \hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}} (see \secref{sec:consts}), and \hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}} (see - \secref{sec:axms-thms}). In particular, type-inference is commonly - available, and result names need not be given. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@nont{\isa{specs}}[] -\rail@endbar -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{decl}}[] -\rail@term{\isa{\isakeyword{where}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mode}{\mbox{\isa{mode}}}}[] -\rail@endbar -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{decl}}[] -\rail@term{\isa{\isakeyword{where}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@end -\rail@begin{4}{\indexdef{}{syntax}{fixes}\hypertarget{syntax.fixes}{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}} -\rail@plus -\rail@bar -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@endbar -\rail@nextplus{3} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{3}{\isa{specs}} -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{2}{\isa{decl}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.axiomatization}{\mbox{\isa{\isacommand{axiomatization}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub m\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} - introduces several constants simultaneously and states axiomatic - properties for these. The constants are marked as being specified - once and for all, which prevents additional specifications being - issued later on. - - Note that axiomatic specifications are only appropriate when - declaring a new logical system; axiomatic specifications are - restricted to global theory contexts. Normal applications should - only use definitional mechanisms! - - \item \hyperlink{command.definition}{\mbox{\isa{\isacommand{definition}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eq{\isaliteral{22}{\isachardoublequote}}} produces an - internal definition \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} according to the specification - given as \isa{eq}, which is then turned into a proven fact. The - given proposition may deviate from internal meta-level equality - according to the rewrite rules declared as \hyperlink{attribute.defn}{\mbox{\isa{defn}}} by the - object-logic. This usually covers object-level equality \isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequote}}} and equivalence \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C6C65667472696768746172726F773E}{\isasymleftrightarrow}}\ B{\isaliteral{22}{\isachardoublequote}}}. End-users normally need not - change the \hyperlink{attribute.defn}{\mbox{\isa{defn}}} setup. - - Definitions may be presented with explicit arguments on the LHS, as - well as additional conditions, e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}f\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ t{\isaliteral{22}{\isachardoublequote}}} instead of - \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ t{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}y\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ {\isadigit{0}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ g\ x\ y\ {\isaliteral{3D}{\isacharequal}}\ u{\isaliteral{22}{\isachardoublequote}}} instead of an - unrestricted \isa{{\isaliteral{22}{\isachardoublequote}}g\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ u{\isaliteral{22}{\isachardoublequote}}}. - - \item \hyperlink{command.abbreviation}{\mbox{\isa{\isacommand{abbreviation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eq{\isaliteral{22}{\isachardoublequote}}} introduces a - syntactic constant which is associated with a certain term according - to the meta-level equality \isa{eq}. - - Abbreviations participate in the usual type-inference process, but - are expanded before the logic ever sees them. Pretty printing of - terms involves higher-order rewriting with rules stemming from - reverted abbreviations. This needs some care to avoid overlapping - or looping syntactic replacements! - - The optional \isa{mode} specification restricts output to a - particular print mode; using ``\isa{input}'' here achieves the - effect of one-way abbreviations. The mode may also include an - ``\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}}'' qualifier that affects the concrete syntax - declared for abbreviations, cf.\ \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} in - \secref{sec:syn-trans}. - - \item \hyperlink{command.print-abbrevs}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}abbrevs}}}} prints all constant abbreviations - of the current context. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Generic declarations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{declaration}\hypertarget{command.declaration}{\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{syntax\_declaration}\hypertarget{command.syntax-declaration}{\hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{declare}\hypertarget{command.declare}{\hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Arbitrary operations on the background context may be wrapped-up as - generic declaration elements. Since the underlying concept of local - theories may be subject to later re-interpretation, there is an - additional dependency on a morphism that tells the difference of the - original declaration context wrt.\ the application context - encountered later on. A fact declaration is an important special - case: it consists of a theorem which is applied to the context by - means of an attribute. - - \begin{railoutput} -\rail@begin{5}{} -\rail@bar -\rail@term{\hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{pervasive}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}~\isa{d} adds the declaration - function \isa{d} of ML type \verb|declaration|, to the current - local theory under construction. In later application contexts, the - function is transformed according to the morphisms being involved in - the interpretation hierarchy. - - If the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}pervasive{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option is given, the corresponding - declaration is applied to all possible contexts involved, including - the global background theory. - - \item \hyperlink{command.syntax-declaration}{\mbox{\isa{\isacommand{syntax{\isaliteral{5F}{\isacharunderscore}}declaration}}}} is similar to \hyperlink{command.declaration}{\mbox{\isa{\isacommand{declaration}}}}, but is meant to affect only ``syntactic'' tools by - convention (such as notation and type-checking information). - - \item \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}}~\isa{thms} declares theorems to the - current local theory context. No theorem binding is involved here, - unlike \hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}} or \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}} (cf.\ - \secref{sec:axms-thms}), so \hyperlink{command.declare}{\mbox{\isa{\isacommand{declare}}}} only has the effect - of applying attributes as included in the theorem specification. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Locales \label{sec:locale}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Locales are parametric named local contexts, consisting of a list of - declaration elements that are modeled after the Isar proof context - commands (cf.\ \secref{sec:proof-context}).% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Locale expressions \label{sec:locale-expr}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A \emph{locale expression} denotes a structured context composed of - instances of existing locales. The context consists of a list of - instances of declaration elements from the locales. Two locale - instances are equal if they are of the same locale and the - parameters are instantiated with equivalent terms. Declaration - elements from equal instances are never repeated, thus avoiding - duplicate declarations. More precisely, declarations from instances - that are subsumed by earlier instances are omitted. - - \begin{railoutput} -\rail@begin{3}{\indexdef{}{syntax}{locale\_expr}\hypertarget{syntax.locale-expr}{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}} -\rail@plus -\rail@nont{\isa{instance}}[] -\rail@nextplus{1} -\rail@cterm{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@endplus -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{2}{\isa{instance}} -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{qualifier}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@nont{\isa{pos{\isaliteral{5F}{\isacharunderscore}}insts}}[] -\rail@nextbar{1} -\rail@nont{\isa{named{\isaliteral{5F}{\isacharunderscore}}insts}}[] -\rail@endbar -\rail@end -\rail@begin{3}{\isa{qualifier}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@bar -\rail@term{\isa{{\isaliteral{3F}{\isacharquery}}}}[] -\rail@nextbar{2} -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@endbar -\rail@endbar -\rail@end -\rail@begin{3}{\isa{pos{\isaliteral{5F}{\isacharunderscore}}insts}} -\rail@plus -\rail@nextplus{1} -\rail@bar -\rail@term{\isa{{\isaliteral{5F}{\isacharunderscore}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@endbar -\rail@endplus -\rail@end -\rail@begin{2}{\isa{named{\isaliteral{5F}{\isacharunderscore}}insts}} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - A locale instance consists of a reference to a locale and either - positional or named parameter instantiations. Identical - instantiations (that is, those that instante a parameter by itself) - may be omitted. The notation `\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5F}{\isacharunderscore}}{\isaliteral{22}{\isachardoublequote}}}' enables to omit the - instantiation for a parameter inside a positional instantiation. - - Terms in instantiations are from the context the locale expressions - is declared in. Local names may be added to this context with the - optional \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} clause. This is useful for shadowing names - bound in outer contexts, and for declaring syntax. In addition, - syntax declarations from one instance are effective when parsing - subsequent instances of the same expression. - - Instances have an optional qualifier which applies to names in - declarations. Names include local definitions and theorem names. - If present, the qualifier itself is either optional - (``\texttt{?}''), which means that it may be omitted on input of the - qualified name, or mandatory (``\texttt{!}''). If neither - ``\texttt{?}'' nor ``\texttt{!}'' are present, the command's default - is used. For \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} and \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}} - the default is ``mandatory'', for \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}} and \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}} the default is ``optional''.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Locale declarations% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{locale}\hypertarget{command.locale}{\hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_locale}\hypertarget{command.print-locale}{\hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_locales}\hypertarget{command.print-locales}{\hyperlink{command.print-locales}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locales}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{method}{intro\_locales}\hypertarget{method.intro-locales}{\hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}}} & : & \isa{method} \\ - \indexdef{}{method}{unfold\_locales}\hypertarget{method.unfold-locales}{\hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}}} & : & \isa{method} \\ - \end{matharray} - - \indexisarelem{fixes}\indexisarelem{constrains}\indexisarelem{assumes} - \indexisarelem{defines}\indexisarelem{notes} - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.locale}{\mbox{\isa{locale}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\rail@begin{5}{\indexdef{}{syntax}{locale}\hypertarget{syntax.locale}{\hyperlink{syntax.locale}{\mbox{\isa{locale}}}}} -\rail@bar -\rail@plus -\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[] -\rail@bar -\rail@nextbar{3} -\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@endbar -\rail@endbar -\rail@end -\rail@begin{12}{\indexdef{}{syntax}{context\_elem}\hypertarget{syntax.context-elem}{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}} -\rail@bar -\rail@term{\isa{\isakeyword{fixes}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.fixes}{\mbox{\isa{fixes}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@nextbar{2} -\rail@term{\isa{\isakeyword{constrains}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@nextplus{3} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{assumes}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.props}{\mbox{\isa{props}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@nextbar{6} -\rail@term{\isa{\isakeyword{defines}}}[] -\rail@plus -\rail@bar -\rail@nextbar{7} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@bar -\rail@nextbar{7} -\rail@nont{\hyperlink{syntax.prop-pat}{\mbox{\isa{prop{\isaliteral{5F}{\isacharunderscore}}pat}}}}[] -\rail@endbar -\rail@nextplus{8} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@nextbar{9} -\rail@term{\isa{\isakeyword{notes}}}[] -\rail@plus -\rail@bar -\rail@nextbar{10} -\rail@nont{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{11} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}loc\ {\isaliteral{3D}{\isacharequal}}\ import\ {\isaliteral{2B}{\isacharplus}}\ body{\isaliteral{22}{\isachardoublequote}}} defines a - new locale \isa{loc} as a context consisting of a certain view of - existing locales (\isa{import}) plus some additional elements - (\isa{body}). Both \isa{import} and \isa{body} are optional; - the degenerate form \hyperlink{command.locale}{\mbox{\isa{\isacommand{locale}}}}~\isa{loc} defines an empty - locale, which may still be useful to collect declarations of facts - later on. Type-inference on locale expressions automatically takes - care of the most general typing that the combined context elements - may acquire. - - The \isa{import} consists of a structured locale expression; see - \secref{sec:proof-context} above. Its for clause defines the local - parameters of the \isa{import}. In addition, locale parameters - whose instantance is omitted automatically extend the (possibly - empty) for clause: they are inserted at its beginning. This means - that these parameters may be referred to from within the expression - and also in the subsequent context elements and provides a - notational convenience for the inheritance of parameters in locale - declarations. - - The \isa{body} consists of context elements. - - \begin{description} - - \item \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\ {\isaliteral{28}{\isacharparenleft}}mx{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} declares a local - parameter of type \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} and mixfix annotation \isa{mx} (both - are optional). The special syntax declaration ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C5354525543545552453E}{\isasymSTRUCTURE}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' means that \isa{x} may be referenced - implicitly in this context. - - \item \hyperlink{element.constrains}{\mbox{\isa{\isakeyword{constrains}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} introduces a type - constraint \isa{{\isaliteral{5C3C7461753E}{\isasymtau}}} on the local parameter \isa{x}. This - element is deprecated. The type constraint should be introduced in - the for clause or the relevant \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} element. - - \item \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} - introduces local premises, similar to \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}} within a - proof (cf.\ \secref{sec:proof-context}). - - \item \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}} defines a previously - declared parameter. This is similar to \hyperlink{command.def}{\mbox{\isa{\isacommand{def}}}} within a - proof (cf.\ \secref{sec:proof-context}), but \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}} - takes an equational proposition instead of variable-term pair. The - left-hand side of the equation may have additional arguments, e.g.\ - ``\hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}''. - - \item \hyperlink{element.notes}{\mbox{\isa{\isakeyword{notes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} - reconsiders facts within a local context. Most notably, this may - include arbitrary declarations in any attribute specifications - included here, e.g.\ a local \hyperlink{attribute.simp}{\mbox{\isa{simp}}} rule. - - The initial \isa{import} specification of a locale expression - maintains a dynamic relation to the locales being referenced - (benefiting from any later fact declarations in the obvious manner). - - \end{description} - - Note that ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C49533E}{\isasymIS}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ p\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}}'' patterns given - in the syntax of \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} and \hyperlink{element.defines}{\mbox{\isa{\isakeyword{defines}}}} above - are illegal in locale definitions. In the long goal format of - \secref{sec:goals}, term bindings may be included as expected, - though. - - \medskip Locale specifications are ``closed up'' by - turning the given text into a predicate definition \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms} and deriving the original assumptions as local lemmas - (modulo local definitions). The predicate statement covers only the - newly specified assumptions, omitting the content of included locale - expressions. The full cumulative view is only provided on export, - involving another predicate \isa{loc} that refers to the complete - specification text. - - In any case, the predicate arguments are those locale parameters - that actually occur in the respective piece of text. Also note that - these predicates operate at the meta-level in theory, but the locale - packages attempts to internalize statements according to the - object-logic setup (e.g.\ replacing \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}} by \isa{{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}}, and - \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} by \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} in HOL; see also - \secref{sec:object-logic}). Separate introduction rules \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro} and \isa{loc{\isaliteral{2E}{\isachardot}}intro} are provided as well. - - \item \hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} prints the - contents of the named locale. The command omits \hyperlink{element.notes}{\mbox{\isa{\isakeyword{notes}}}} - elements by default. Use \hyperlink{command.print-locale}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locale}}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} to - have them included. - - \item \hyperlink{command.print-locales}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}locales}}}} prints the names of all locales - of the current theory. - - \item \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}} and \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}} - repeatedly expand all introduction rules of locale predicates of the - theory. While \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}} only applies the \isa{loc{\isaliteral{2E}{\isachardot}}intro} introduction rules and therefore does not decend to - assumptions, \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}} is more aggressive and applies - \isa{loc{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro} as well. Both methods are aware of locale - specifications entailed by the context, both from target statements, - and from interpretations (see below). New goals that are entailed - by the current context are discharged automatically. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Locale interpretation% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{interpretation}\hypertarget{command.interpretation}{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{interpret}\hypertarget{command.interpret}{\hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof{\isaliteral{28}{\isacharparenleft}}state{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{7C}{\isacharbar}}\ proof{\isaliteral{28}{\isacharparenleft}}chain{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{sublocale}\hypertarget{command.sublocale}{\hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_dependencies}\hypertarget{command.print-dependencies}{\hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_interps}\hypertarget{command.print-interps}{\hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Locale expressions may be instantiated, and the instantiated facts - added to the current context. This requires a proof of the - instantiated specification and is called \emph{locale - interpretation}. Interpretation is possible in locales (command - \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}), theories (command \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}) and also within proof bodies (command \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}). - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}}[] -\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{equations}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}}[] -\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{equations}}[] -\rail@endbar -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[] -\rail@cr{3} -\rail@bar -\rail@nextbar{4} -\rail@nont{\isa{equations}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{21}{\isacharbang}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.locale-expr}{\mbox{\isa{locale{\isaliteral{5F}{\isacharunderscore}}expr}}}}[] -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}}[] -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\rail@begin{3}{\isa{equations}} -\rail@term{\isa{\isakeyword{where}}}[] -\rail@plus -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}} - interprets \isa{expr} in the theory. The command generates proof - obligations for the instantiated specifications (assumes and defines - elements). Once these are discharged by the user, instantiated - facts are added to the theory in a post-processing phase. - - Free variables in the interpreted expression are allowed. They are - turned into schematic variables in the generated declarations. In - order to use a free variable whose name is already bound in the - context --- for example, because a constant of that name exists --- - it may be added to the \hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}} clause. - - Additional equations, which are unfolded during - post-processing, may be given after the keyword \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}}. - This is useful for interpreting concepts introduced through - definitions. The equations must be proved. - - The command is aware of interpretations already active in the - theory, but does not simplify the goal automatically. In order to - simplify the proof obligations use methods \hyperlink{method.intro-locales}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}locales}}} - or \hyperlink{method.unfold-locales}{\mbox{\isa{unfold{\isaliteral{5F}{\isacharunderscore}}locales}}}. Post-processing is not applied to - facts of interpretations that are already active. This avoids - duplication of interpreted facts, in particular. Note that, in the - case of a locale with import, parts of the interpretation may - already be active. The command will only process facts for new - parts. - - Adding facts to locales has the effect of adding interpreted facts - to the theory for all interpretations as well. That is, - interpretations dynamically participate in any facts added to - locales. Note that if a theory inherits additional facts for a - locale through one parent and an interpretation of that locale - through another parent, the additional facts will not be - interpreted. - - \item \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}} interprets - \isa{expr} in the proof context and is otherwise similar to - interpretation in theories. Note that rewrite rules given to - \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}} after the \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} keyword should be - explicitly universally quantified. - - \item \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ expr\ {\isaliteral{5C3C57484552453E}{\isasymWHERE}}\ eqns{\isaliteral{22}{\isachardoublequote}}} - interprets \isa{expr} in the locale \isa{name}. A proof that - the specification of \isa{name} implies the specification of - \isa{expr} is required. As in the localized version of the - theorem command, the proof is in the context of \isa{name}. After - the proof obligation has been discharged, the facts of \isa{expr} - become part of locale \isa{name} as \emph{derived} context - elements and are available when the context \isa{name} is - subsequently entered. Note that, like import, this is dynamic: - facts added to a locale part of \isa{expr} after interpretation - become also available in \isa{name}. - - Only specification fragments of \isa{expr} that are not already - part of \isa{name} (be it imported, derived or a derived fragment - of the import) are considered in this process. This enables - circular interpretations provided that no infinite chains are - generated in the locale hierarchy. - - If interpretations of \isa{name} exist in the current theory, the - command adds interpretations for \isa{expr} as well, with the same - qualifier, although only for fragments of \isa{expr} that are not - interpreted in the theory already. - - Equations given after \hyperlink{keyword.where}{\mbox{\isa{\isakeyword{where}}}} amend the morphism through - which \isa{expr} is interpreted. This enables to map definitions - from the interpreted locales to entities of \isa{name}. This - feature is experimental. - - \item \hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}expr{\isaliteral{22}{\isachardoublequote}}} is useful for - understanding the effect of an interpretation of \isa{{\isaliteral{22}{\isachardoublequote}}expr{\isaliteral{22}{\isachardoublequote}}}. It - lists all locale instances for which interpretations would be added - to the current context. Variant \hyperlink{command.print-dependencies}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}dependencies}}}}\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{21}{\isacharbang}}{\isaliteral{22}{\isachardoublequote}}} prints all locale instances that - would be considered for interpretation, and would be interpreted in - an empty context (that is, without interpretations). - - \item \hyperlink{command.print-interps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}interps}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} lists all - interpretations of \isa{{\isaliteral{22}{\isachardoublequote}}locale{\isaliteral{22}{\isachardoublequote}}} in the current theory or proof - context, including those due to a combination of a \hyperlink{command.interpretation}{\mbox{\isa{\isacommand{interpretation}}}} or \hyperlink{command.interpret}{\mbox{\isa{\isacommand{interpret}}}} and one or several - \hyperlink{command.sublocale}{\mbox{\isa{\isacommand{sublocale}}}} declarations. - - \end{description} - - \begin{warn} - Since attributes are applied to interpreted theorems, - interpretation may modify the context of common proof tools, e.g.\ - the Simplifier or Classical Reasoner. As the behavior of such - tools is \emph{not} stable under interpretation morphisms, manual - declarations might have to be added to the target context of the - interpretation to revert such declarations. - \end{warn} - - \begin{warn} - An interpretation in a theory or proof context may subsume previous - interpretations. This happens if the same specification fragment - is interpreted twice and the instantiation of the second - interpretation is more general than the interpretation of the - first. The locale package does not attempt to remove subsumed - interpretations. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Classes \label{sec:class}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{class}\hypertarget{command.class}{\hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{instantiation}\hypertarget{command.instantiation}{\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{instance}\hypertarget{command.instance}{\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{28}{\isacharparenleft}}prove{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{subclass}\hypertarget{command.subclass}{\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{print\_classes}\hypertarget{command.print-classes}{\hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}classes}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{class\_deps}\hypertarget{command.class-deps}{\hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isaliteral{5F}{\isacharunderscore}}deps}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{method}{intro\_classes}\hypertarget{method.intro-classes}{\hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}}} & : & \isa{method} \\ - \end{matharray} - - A class is a particular locale with \emph{exactly one} type variable - \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}}. Beyond the underlying locale, a corresponding type class - is established which is interpreted logically as axiomatic type - class \cite{Wenzel:1997:TPHOL} whose logical content are the - assumptions of the locale. Thus, classes provide the full - generality of locales combined with the commodity of type classes - (notably type-inference). See \cite{isabelle-classes} for a short - tutorial. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}}[] -\rail@nont{\isa{class{\isaliteral{5F}{\isacharunderscore}}spec}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@endbar -\rail@end -\rail@begin{5}{\isa{class{\isaliteral{5F}{\isacharunderscore}}spec}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@bar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{2B}{\isacharplus}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@nextbar{2} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextbar{3} -\rail@plus -\rail@nont{\hyperlink{syntax.context-elem}{\mbox{\isa{context{\isaliteral{5F}{\isacharunderscore}}elem}}}}[] -\rail@nextplus{4} -\rail@endplus -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{1} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[] -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{5}{} -\rail@term{\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[] -\rail@nextbar{4} -\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.class}{\mbox{\isa{\isacommand{class}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3D}{\isacharequal}}\ superclasses\ {\isaliteral{2B}{\isacharplus}}\ body{\isaliteral{22}{\isachardoublequote}}} defines - a new class \isa{c}, inheriting from \isa{superclasses}. This - introduces a locale \isa{c} with import of all locales \isa{superclasses}. - - Any \hyperlink{element.fixes}{\mbox{\isa{\isakeyword{fixes}}}} in \isa{body} are lifted to the global - theory level (\emph{class operations} \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ f\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} of class \isa{c}), mapping the local type parameter - \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}} to a schematic type variable \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{22}{\isachardoublequote}}}. - - Likewise, \hyperlink{element.assumes}{\mbox{\isa{\isakeyword{assumes}}}} in \isa{body} are also lifted, - mapping each local parameter \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} to its - corresponding global constant \isa{{\isaliteral{22}{\isachardoublequote}}f\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}. The - corresponding introduction rule is provided as \isa{c{\isaliteral{5F}{\isacharunderscore}}class{\isaliteral{5F}{\isacharunderscore}}axioms{\isaliteral{2E}{\isachardot}}intro}. This rule should be rarely needed directly - --- the \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}} method takes care of the details of - class membership proofs. - - \item \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} opens a theory target (cf.\ \secref{sec:target}) which - allows to specify class operations \isa{{\isaliteral{22}{\isachardoublequote}}f\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ f\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} corresponding - to sort \isa{s} at the particular type instance \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}}. A plain \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} command in the - target body poses a goal stating these type arities. The target is - concluded by an \indexref{local}{command}{end}\hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}} command. - - Note that a list of simultaneous type constructors may be given; - this corresponds nicely to mutually recursive type definitions, e.g.\ - in Isabelle/HOL. - - \item \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} in an instantiation target body sets - up a goal stating the type arities claimed at the opening \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}}. The proof would usually proceed by \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}}, and then establish the characteristic theorems of - the type classes involved. After finishing the proof, the - background theory will be augmented by the proven type arities. - - On the theory level, \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}} provides a convenient way to instantiate a type class with no - need to specify operations: one can continue with the - instantiation proof immediately. - - \item \hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}}~\isa{c} in a class context for class - \isa{d} sets up a goal stating that class \isa{c} is logically - contained in class \isa{d}. After finishing the proof, class - \isa{d} is proven to be subclass \isa{c} and the locale \isa{c} is interpreted into \isa{d} simultaneously. - - A weakend form of this is available through a further variant of - \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}: \hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} opens - a proof that class \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} implies \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E697375623E}{}\isactrlisub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} without reference - to the underlying locales; this is useful if the properties to prove - the logical connection are not sufficent on the locale level but on - the theory level. - - \item \hyperlink{command.print-classes}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}classes}}}} prints all classes in the current - theory. - - \item \hyperlink{command.class-deps}{\mbox{\isa{\isacommand{class{\isaliteral{5F}{\isacharunderscore}}deps}}}} visualizes all classes and their - subclass relations as a Hasse diagram. - - \item \hyperlink{method.intro-classes}{\mbox{\isa{intro{\isaliteral{5F}{\isacharunderscore}}classes}}} repeatedly expands all class - introduction rules of this theory. Note that this method usually - needs not be named explicitly, as it is already included in the - default proof step (e.g.\ of \hyperlink{command.proof}{\mbox{\isa{\isacommand{proof}}}}). In particular, - instantiation of trivial (syntactic) classes may be performed by a - single ``\hyperlink{command.ddot}{\mbox{\isa{\isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}}}}'' proof step. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{The class target% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -%FIXME check - - A named context may refer to a locale (cf.\ \secref{sec:target}). - If this locale is also a class \isa{c}, apart from the common - locale target behaviour the following happens. - - \begin{itemize} - - \item Local constant declarations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} referring to the - local type parameter \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}} and local parameters \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} - are accompanied by theory-level constants \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} - referring to theory-level class operations \isa{{\isaliteral{22}{\isachardoublequote}}f{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}}. - - \item Local theorem bindings are lifted as are assumptions. - - \item Local syntax refers to local operations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} and - global operations \isa{{\isaliteral{22}{\isachardoublequote}}g{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ c{\isaliteral{5D}{\isacharbrackright}}{\isaliteral{22}{\isachardoublequote}}} uniformly. Type inference - resolves ambiguities. In rare cases, manual type annotations are - needed. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Co-regularity of type classes and arities% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The class relation together with the collection of - type-constructor arities must obey the principle of - \emph{co-regularity} as defined below. - - \medskip For the subsequent formulation of co-regularity we assume - that the class relation is closed by transitivity and reflexivity. - Moreover the collection of arities \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{22}{\isachardoublequote}}} is - completed such that \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} - implies \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s{\isaliteral{29}{\isacharparenright}}c{\isaliteral{27}{\isacharprime}}{\isaliteral{22}{\isachardoublequote}}} for all such declarations. - - Treating sorts as finite sets of classes (meaning the intersection), - the class relation \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} is extended to sorts as - follows: - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{2E}{\isachardot}}\ {\isaliteral{5C3C6578697374733E}{\isasymexists}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C696E3E}{\isasymin}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2E}{\isachardot}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} - \] - - This relation on sorts is further extended to tuples of sorts (of - the same length) in the component-wise way. - - \smallskip Co-regularity of the class relation together with the - arities relation means: - \[ - \isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}\isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ \isaliteral{5C3C5E7665633E}{}\isactrlvec s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} - \] - \noindent for all such arities. In other words, whenever the result - classes of some type-constructor arities are related, then the - argument sorts need to be related in the same way. - - \medskip Co-regularity is a very fundamental property of the - order-sorted algebra of types. For example, it entails principle - types and most general unifiers, e.g.\ see \cite{nipkow-prehofer}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Unrestricted overloading% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{overloading}\hypertarget{command.overloading}{\hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Isabelle/Pure's definitional schemes support certain forms of - overloading (see \secref{sec:consts}). Overloading means that a - constant being declared as \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ decl{\isaliteral{22}{\isachardoublequote}}} may be - defined separately on type instances - \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ decl{\isaliteral{22}{\isachardoublequote}}} - for each type constructor \isa{t}. At most occassions - overloading will be used in a Haskell-like fashion together with - type classes by means of \hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} (see - \secref{sec:class}). Sometimes low-level overloading is desirable. - The \hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}} target provides a convenient view for - end-users. - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}}[] -\rail@plus -\rail@nont{\isa{spec}}[] -\rail@nextplus{1} -\rail@endplus -\rail@term{\isa{\isakeyword{begin}}}[] -\rail@end -\rail@begin{2}{\isa{spec}} -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5C3C65717569763E}{\isasymequiv}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{unchecked}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.overloading}{\mbox{\isa{\isacommand{overloading}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C414E443E}{\isasymAND}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C424547494E3E}{\isasymBEGIN}}{\isaliteral{22}{\isachardoublequote}}} - opens a theory target (cf.\ \secref{sec:target}) which allows to - specify constants with overloaded definitions. These are identified - by an explicitly given mapping from variable names \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} to - constants \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} at particular type instances. The - definitions themselves are established using common specification - tools, using the names \isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub i{\isaliteral{22}{\isachardoublequote}}} as reference to the - corresponding constants. The target is concluded by \hyperlink{command.local.end}{\mbox{\isa{\isacommand{end}}}}. - - A \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}unchecked{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option disables global dependency checks for - the corresponding definition, which is occasionally useful for - exotic overloading (see \secref{sec:consts} for a precise description). - It is at the discretion of the user to avoid - malformed theory specifications! - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Incorporating ML code \label{sec:ML}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{ML\_file}\hypertarget{command.ML-file}{\hyperlink{command.ML-file}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}file}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{ML}\hypertarget{command.ML}{\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{ML\_prf}\hypertarget{command.ML-prf}{\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}proof\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ proof{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{ML\_val}\hypertarget{command.ML-val}{\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{ML\_command}\hypertarget{command.ML-command}{\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{setup}\hypertarget{command.setup}{\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{local\_setup}\hypertarget{command.local-setup}{\hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{attribute\_setup}\hypertarget{command.attribute-setup}{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{command.ML-file}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}file}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@end -\rail@begin{6}{} -\rail@bar -\rail@term{\hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}}}[] -\rail@nextbar{2} -\rail@term{\hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}}}[] -\rail@nextbar{3} -\rail@term{\hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}}}[] -\rail@nextbar{4} -\rail@term{\hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}}[] -\rail@nextbar{5} -\rail@term{\hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.ML-file}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}file}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{22}{\isachardoublequote}}} reads and evaluates the - given ML file. The current theory context is passed down to the ML - toplevel and may be modified, using \verb|Context.>>| or derived ML - commands. Top-level ML bindings are stored within the (global or - local) theory context. - - \item \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} is similar to \hyperlink{command.ML-file}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}file}}}}, but evaluates directly the given \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}}. - Top-level ML bindings are stored within the (global or local) theory - context. - - \item \hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}} is analogous to \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}} but works - within a proof context. Top-level ML bindings are stored within the - proof context in a purely sequential fashion, disregarding the - nested proof structure. ML bindings introduced by \hyperlink{command.ML-prf}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}prf}}}} are discarded at the end of the proof. - - \item \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}} and \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}} are diagnostic - versions of \hyperlink{command.ML}{\mbox{\isa{\isacommand{ML}}}}, which means that the context may not be - updated. \hyperlink{command.ML-val}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}val}}}} echos the bindings produced at the ML - toplevel, but \hyperlink{command.ML-command}{\mbox{\isa{\isacommand{ML{\isaliteral{5F}{\isacharunderscore}}command}}}} is silent. - - \item \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} changes the current theory - context by applying \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}}, which refers to an ML expression - of type \verb|theory -> theory|. This enables to initialize - any object-logic specific tools and packages written in ML, for - example. - - \item \hyperlink{command.local-setup}{\mbox{\isa{\isacommand{local{\isaliteral{5F}{\isacharunderscore}}setup}}}} is similar to \hyperlink{command.setup}{\mbox{\isa{\isacommand{setup}}}} for - a local theory context, and an ML expression of type \verb|local_theory -> local_theory|. This allows to - invoke local theory specification packages without going through - concrete outer syntax, for example. - - \item \hyperlink{command.attribute-setup}{\mbox{\isa{\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ text\ description{\isaliteral{22}{\isachardoublequote}}} - defines an attribute in the current theory. The given \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} has to be an ML expression of type - \verb|attribute context_parser|, cf.\ basic parsers defined in - structure \verb|Args| and \verb|Attrib|. - - In principle, attributes can operate both on a given theorem and the - implicit context, although in practice only one is modified and the - other serves as parameter. Here are examples for these two cases: - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimML -\ \ % -\endisadelimML -% -\isatagML -\isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ my{\isaliteral{5F}{\isacharunderscore}}rule\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ \ \ Attrib{\isaliteral{2E}{\isachardot}}thms\ {\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}fn\ ths\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}rule{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline -\ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ \ \ \ \ let\ val\ th{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ th\ OF\ ths\isanewline -\ \ \ \ \ \ \ \ \ \ in\ th{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}\isanewline -\isanewline -\ \ \isacommand{attribute{\isaliteral{5F}{\isacharunderscore}}setup}\isamarkupfalse% -\ my{\isaliteral{5F}{\isacharunderscore}}declaration\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline -\ \ \ \ Attrib{\isaliteral{2E}{\isachardot}}thms\ {\isaliteral{3E}{\isachargreater}}{\isaliteral{3E}{\isachargreater}}\ {\isaliteral{28}{\isacharparenleft}}fn\ ths\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ Thm{\isaliteral{2E}{\isachardot}}declaration{\isaliteral{5F}{\isacharunderscore}}attribute\isanewline -\ \ \ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}fn\ th{\isaliteral{3A}{\isacharcolon}}\ thm\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\ fn\ context{\isaliteral{3A}{\isacharcolon}}\ Context{\isaliteral{2E}{\isachardot}}generic\ {\isaliteral{3D}{\isacharequal}}{\isaliteral{3E}{\isachargreater}}\isanewline -\ \ \ \ \ \ \ \ \ \ let\ val\ context{\isaliteral{27}{\isacharprime}}\ {\isaliteral{3D}{\isacharequal}}\ context\isanewline -\ \ \ \ \ \ \ \ \ \ in\ context{\isaliteral{27}{\isacharprime}}\ end{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2A7D}{\isacharverbatimclose}}% -\endisatagML -{\isafoldML}% -% -\isadelimML -% -\endisadelimML -% -\isamarkupsection{Primitive specification elements% -} -\isamarkuptrue% -% -\isamarkupsubsection{Type classes and sorts \label{sec:classes}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{classes}\hypertarget{command.classes}{\hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{classrel}\hypertarget{command.classrel}{\hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\ - \indexdef{}{command}{default\_sort}\hypertarget{command.default-sort}{\hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.classdecl}{\mbox{\isa{classdecl}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{3}{} -\rail@term{\hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@bar -\rail@term{\isa{{\isaliteral{3C}{\isacharless}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{2} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@end -\rail@begin{1}{} -\rail@term{\hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}}[] -\rail@nont{\hyperlink{syntax.sort}{\mbox{\isa{sort}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.classes}{\mbox{\isa{\isacommand{classes}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}} declares class - \isa{c} to be a subclass of existing classes \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. - Isabelle implicitly maintains the transitive closure of the class - hierarchy. Cyclic class structures are not permitted. - - \item \hyperlink{command.classrel}{\mbox{\isa{\isacommand{classrel}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}\ c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} states subclass - relations between existing classes \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}c\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}}. - This is done axiomatically! The \indexref{}{command}{subclass}\hyperlink{command.subclass}{\mbox{\isa{\isacommand{subclass}}}} and - \indexref{}{command}{instance}\hyperlink{command.instance}{\mbox{\isa{\isacommand{instance}}}} commands (see \secref{sec:class}) provide - a way to introduce proven class relations. - - \item \hyperlink{command.default-sort}{\mbox{\isa{\isacommand{default{\isaliteral{5F}{\isacharunderscore}}sort}}}}~\isa{s} makes sort \isa{s} the - new default sort for any type variable that is given explicitly in - the text, but lacks a sort constraint (wrt.\ the current context). - Type variables generated by type inference are not affected. - - Usually the default sort is only changed when defining a new - object-logic. For example, the default sort in Isabelle/HOL is - \isa{type}, the class of all HOL types. - - When merging theories, the default sorts of the parents are - logically intersected, i.e.\ the representations as lists of classes - are joined. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Types and type abbreviations \label{sec:types-pure}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{type\_synonym}\hypertarget{command.type-synonym}{\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{typedecl}\hypertarget{command.typedecl}{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{arities}\hypertarget{command.arities}{\hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}}[] -\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}}[] -\rail@nont{\hyperlink{syntax.typespec}{\mbox{\isa{typespec}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.arity}{\mbox{\isa{arity}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.type-synonym}{\mbox{\isa{\isacommand{type{\isaliteral{5F}{\isacharunderscore}}synonym}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}} - introduces a \emph{type synonym} \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} for the - existing type \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7461753E}{\isasymtau}}{\isaliteral{22}{\isachardoublequote}}}. Unlike actual type definitions, as are - available in Isabelle/HOL for example, type synonyms are merely - syntactic abbreviations without any logical significance. - Internally, type synonyms are fully expanded. - - \item \hyperlink{command.typedecl}{\mbox{\isa{\isacommand{typedecl}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}\ t{\isaliteral{22}{\isachardoublequote}}} declares a new - type constructor \isa{t}. If the object-logic defines a base sort - \isa{s}, then the constructor is declared to operate on that, via - the axiomatic specification \hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}}. - - \item \hyperlink{command.arities}{\mbox{\isa{\isacommand{arities}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}t\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{28}{\isacharparenleft}}s\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ s\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{29}{\isacharparenright}}s{\isaliteral{22}{\isachardoublequote}}} augments - Isabelle's order-sorted signature of types by new type constructor - arities. This is done axiomatically! The \indexref{}{command}{instantiation}\hyperlink{command.instantiation}{\mbox{\isa{\isacommand{instantiation}}}} - target (see \secref{sec:class}) provides a way to introduce - proven type arities. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Constants and definitions \label{sec:consts}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{consts}\hypertarget{command.consts}{\hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{defs}\hypertarget{command.defs}{\hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - Definitions essentially express abbreviations within the logic. The - simplest form of a definition is \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}, where \isa{c} is a newly declared constant. Isabelle also allows derived forms - where the arguments of \isa{c} appear on the left, abbreviating a - prefix of \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-abstractions, e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x\ y{\isaliteral{2E}{\isachardot}}\ t{\isaliteral{22}{\isachardoublequote}}} may be - written more conveniently as \isa{{\isaliteral{22}{\isachardoublequote}}c\ x\ y\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}. Moreover, - definitions may be weakened by adding arbitrary pre-conditions: - \isa{{\isaliteral{22}{\isachardoublequote}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ c\ x\ y\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t{\isaliteral{22}{\isachardoublequote}}}. - - \medskip The built-in well-formedness conditions for definitional - specifications are: - - \begin{itemize} - - \item Arguments (on the left-hand side) must be distinct variables. - - \item All variables on the right-hand side must also appear on the - left-hand side. - - \item All type variables on the right-hand side must also appear on - the left-hand side; this prohibits \isa{{\isaliteral{22}{\isachardoublequote}}{\isadigit{0}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ length\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ list{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} for example. - - \item The definition must not be recursive. Most object-logics - provide definitional principles that can be used to express - recursion safely. - - \end{itemize} - - The right-hand side of overloaded definitions may mention overloaded constants - recursively at type instances corresponding to the immediate - argument types \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}. Incomplete - specification patterns impose global constraints on all occurrences, - e.g.\ \isa{{\isaliteral{22}{\isachardoublequote}}d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}} on the left-hand side means that all - corresponding occurrences on some right-hand side need to be an - instance of this, general \isa{{\isaliteral{22}{\isachardoublequote}}d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C74696D65733E}{\isasymtimes}}\ {\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{22}{\isachardoublequote}}} will be disallowed. - - \begin{railoutput} -\rail@begin{3}{} -\rail@term{\hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}}}[] -\rail@nont{\hyperlink{syntax.type}{\mbox{\isa{type}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@endbar -\rail@nextplus{2} -\rail@endplus -\rail@end -\rail@begin{2}{} -\rail@term{\hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@nont{\isa{opt}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}[] -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{2}{\isa{opt}} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{unchecked}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{\isakeyword{overloaded}}}[] -\rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}c\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7369676D613E}{\isasymsigma}}{\isaliteral{22}{\isachardoublequote}}} declares constant \isa{c} to have any instance of type scheme \isa{{\isaliteral{5C3C7369676D613E}{\isasymsigma}}}. The optional - mixfix annotations may attach concrete syntax to the constants - declared. - - \item \hyperlink{command.defs}{\mbox{\isa{\isacommand{defs}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name{\isaliteral{3A}{\isacharcolon}}\ eqn{\isaliteral{22}{\isachardoublequote}}} introduces \isa{eqn} - as a definitional axiom for some existing constant. - - The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}unchecked{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option disables global dependency checks - for this definition, which is occasionally useful for exotic - overloading. It is at the discretion of the user to avoid malformed - theory specifications! - - The \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}overloaded{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} option declares definitions to be - potentially overloaded. Unless this option is given, a warning - message would be issued for any definitional equation with a more - special type than that of the corresponding constant declaration. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Axioms and theorems \label{sec:axms-thms}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{axioms}\hypertarget{command.axioms}{\hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\ - \indexdef{}{command}{lemmas}\hypertarget{command.lemmas}{\hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{theorems}\hypertarget{command.theorems}{\hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}local{\isaliteral{5F}{\isacharunderscore}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ local{\isaliteral{5F}{\isacharunderscore}}theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{2}{} -\rail@term{\hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.axmdecl}{\mbox{\isa{axmdecl}}}}[] -\rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\rail@begin{6}{} -\rail@bar -\rail@term{\hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}}[] -\rail@nextbar{1} -\rail@term{\hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@nont{\hyperlink{syntax.target}{\mbox{\isa{target}}}}[] -\rail@endbar -\rail@cr{3} -\rail@plus -\rail@bar -\rail@nextbar{4} -\rail@nont{\hyperlink{syntax.thmdef}{\mbox{\isa{thmdef}}}}[] -\rail@endbar -\rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@bar -\rail@nextbar{4} -\rail@term{\isa{\isakeyword{for}}}[] -\rail@plus -\rail@nont{\hyperlink{syntax.vars}{\mbox{\isa{vars}}}}[] -\rail@nextplus{5} -\rail@cterm{\isa{\isakeyword{and}}}[] -\rail@endplus -\rail@endbar -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.axioms}{\mbox{\isa{\isacommand{axioms}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C7068693E}{\isasymphi}}{\isaliteral{22}{\isachardoublequote}}} introduces arbitrary - statements as axioms of the meta-logic. In fact, axioms are - ``axiomatic theorems'', and may be referred later just as any other - theorem. - - Axioms are usually only introduced when declaring new logical - systems. Everyday work is typically done the hard way, with proper - definitions and proven theorems. - - \item \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}a\ {\isaliteral{3D}{\isacharequal}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ b\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{22}{\isachardoublequote}}}~\indexdef{}{keyword}{for}\hypertarget{keyword.for}{\hyperlink{keyword.for}{\mbox{\isa{\isakeyword{for}}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub m{\isaliteral{22}{\isachardoublequote}}} evaluates given facts (with attributes) in - the current context, which may be augmented by local variables. - Results are standardized before being stored, i.e.\ schematic - variables are renamed to enforce index \isa{{\isaliteral{22}{\isachardoublequote}}{\isadigit{0}}{\isaliteral{22}{\isachardoublequote}}} uniformly. - - \item \hyperlink{command.theorems}{\mbox{\isa{\isacommand{theorems}}}} is the same as \hyperlink{command.lemmas}{\mbox{\isa{\isacommand{lemmas}}}}, but - marks the result as a different kind of facts. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Oracles% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcll} - \indexdef{}{command}{oracle}\hypertarget{command.oracle}{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} & (axiomatic!) \\ - \end{matharray} - - Oracles allow Isabelle to take advantage of external reasoners such - as arithmetic decision procedures, model checkers, fast tautology - checkers or computer algebra systems. Invoked as an oracle, an - external reasoner can create arbitrary Isabelle theorems. - - It is the responsibility of the user to ensure that the external - reasoner is as trustworthy as the application requires. Another - typical source of errors is the linkup between Isabelle and the - external tool, not just its concrete implementation, but also the - required translation between two different logical environments. - - Isabelle merely guarantees well-formedness of the propositions being - asserted, and records within the internal derivation object how - presumed theorems depend on unproven suppositions. - - \begin{railoutput} -\rail@begin{1}{} -\rail@term{\hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}}[] -\rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[] -\rail@nont{\hyperlink{syntax.text}{\mbox{\isa{text}}}}[] -\rail@end -\end{railoutput} - - - \begin{description} - - \item \hyperlink{command.oracle}{\mbox{\isa{\isacommand{oracle}}}}~\isa{{\isaliteral{22}{\isachardoublequote}}name\ {\isaliteral{3D}{\isacharequal}}\ text{\isaliteral{22}{\isachardoublequote}}} turns the given ML - expression \isa{{\isaliteral{22}{\isachardoublequote}}text{\isaliteral{22}{\isachardoublequote}}} of type \verb|'a -> cterm| into an - ML function of type \verb|'a -> thm|, which is bound to the - global identifier \verb|name|. This acts like an infinitary - specification of axioms! Invoking the oracle only works within the - scope of the resulting theory. - - \end{description} - - See \verb|~~/src/HOL/ex/Iff_Oracle.thy| for a worked example of - defining a new primitive rule as oracle, and turning it into a proof - method.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{Name spaces% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{hide\_class}\hypertarget{command.hide-class}{\hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{hide\_type}\hypertarget{command.hide-type}{\hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{hide\_const}\hypertarget{command.hide-const}{\hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{hide\_fact}\hypertarget{command.hide-fact}{\hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - \begin{railoutput} -\rail@begin{4}{} -\rail@bar -\rail@nont{\hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}}[] -\rail@nextbar{1} -\rail@nont{\hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}}[] -\rail@nextbar{2} -\rail@nont{\hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}}[] -\rail@nextbar{3} -\rail@nont{\hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}}}[] -\rail@endbar -\rail@bar -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{open}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar -\rail@plus -\rail@nont{\hyperlink{syntax.nameref}{\mbox{\isa{nameref}}}}[] -\rail@nextplus{1} -\rail@endplus -\rail@end -\end{railoutput} - - - Isabelle organizes any kind of name declarations (of types, - constants, theorems etc.) by separate hierarchically structured name - spaces. Normally the user does not have to control the behavior of - name spaces by hand, yet the following commands provide some way to - do so. - - \begin{description} - - \item \hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}~\isa{names} fully removes class - declarations from a given name space; with the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}open{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} - option, only the base name is hidden. - - Note that hiding name space accesses has no impact on logical - declarations --- they remain valid internally. Entities that are no - longer accessible to the user are printed with the special qualifier - ``\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{3F}{\isacharquery}}{\isaliteral{22}{\isachardoublequote}}}'' prefixed to the full internal name. - - \item \hyperlink{command.hide-type}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}type}}}}, \hyperlink{command.hide-const}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}const}}}}, and \hyperlink{command.hide-fact}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}fact}}}} are similar to \hyperlink{command.hide-class}{\mbox{\isa{\isacommand{hide{\isaliteral{5F}{\isacharunderscore}}class}}}}, but hide types, - constants, and facts, respectively. - - \end{description}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\isanewline -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Symbols.tex --- a/doc-src/IsarRef/Thy/document/Symbols.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,72 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Symbols}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Symbols\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Predefined Isabelle symbols \label{app:symbols}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle supports an infinite number of non-ASCII symbols, which are - represented in source text as \verb|\|\verb|<|\isa{name}\verb|>| (where \isa{name} may be any identifier). It - is left to front-end tools how to present these symbols to the user. - The collection of predefined standard symbols given below is - available by default for Isabelle document output, due to - appropriate definitions of \verb|\|\verb|isasym|\isa{name} for each \verb|\|\verb|<|\isa{name}\verb|>| in the \verb|isabellesym.sty| file. Most of these symbols - are displayed properly in Proof~General and Isabelle/jEdit. - - Moreover, any single symbol (or ASCII character) may be prefixed by - \verb|\|\verb|<^sup>|, for superscript and \verb|\|\verb|<^sub>|, for subscript, such as \verb|A\|\verb|<^sup>\|, for \isa{{\isaliteral{22}{\isachardoublequote}}A\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{5C3C737461723E}{\isasymstar}}{\isaliteral{22}{\isachardoublequote}}} the alternative - versions \verb|\|\verb|<^isub>| and \verb|\|\verb|<^isup>| are considered as quasi letters and may - be used within identifiers. Sub- and superscripts that span a - region of text are marked up with \verb|\|\verb|<^bsub>|\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<^esub>|, and - \verb|\|\verb|<^bsup>|\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}}\verb|\|\verb|<^esup>| respectively. Furthermore, all ASCII - characters and most other symbols may be printed in bold by - prefixing \verb|\|\verb|<^bold>| such as \verb|\|\verb|<^bold>\|\verb|| for \isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E626F6C643E}{}\isactrlbold {\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{22}{\isachardoublequote}}}. Note that \verb|\|\verb|<^bold>|, may - \emph{not} be combined with sub- or superscripts for single symbols. - - Further details of Isabelle document preparation are covered in - \chref{ch:document-prep}. - - \begin{center} - \begin{isabellebody} - \input{syms} - \end{isabellebody} - \end{center}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/Synopsis.tex --- a/doc-src/IsarRef/Thy/document/Synopsis.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2718 +0,0 @@ -% -\begin{isabellebody}% -\def\isabellecontext{Synopsis}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isatagtheory -\isacommand{theory}\isamarkupfalse% -\ Synopsis\isanewline -\isakeyword{imports}\ Base\ Main\isanewline -\isakeyword{begin}% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -% -\endisadelimtheory -% -\isamarkupchapter{Synopsis% -} -\isamarkuptrue% -% -\isamarkupsection{Notepad% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -An Isar proof body serves as mathematical notepad to compose logical - content, consisting of types, terms, facts.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Types and terms% -} -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Locally fixed entities:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ \ \ % -\isamarkupcmt{local constant, without any type information yet% -} -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ \ % -\isamarkupcmt{variant with explicit type-constraint for subsequent use% -} -\isanewline -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ a\ b\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \ % -\isamarkupcmt{type assignment at first occurrence in concrete term% -} -% -\begin{isamarkuptxt}% -Definitions (non-polymorphic):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{def}\isamarkupfalse% -\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{22}{\isachardoublequoteopen}}t{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{27}{\isacharprime}}a{\isaliteral{22}{\isachardoublequoteclose}}% -\begin{isamarkuptxt}% -Abbreviations (polymorphic):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{let}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}x{\isaliteral{2E}{\isachardot}}\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\ \ \isacommand{term}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{3F}{\isacharquery}}f\ {\isaliteral{3F}{\isacharquery}}f{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Notation:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{write}\isamarkupfalse% -\ x\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Facts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A fact is a simultaneous list of theorems.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsubsection{Producing facts% -} -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Via assumption (``lambda''):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A% -\begin{isamarkuptxt}% -Via proof (``let''):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ b{\isaliteral{3A}{\isacharcolon}}\ B\ \isacommand{sorry}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Via abbreviation (``let''):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{note}\isamarkupfalse% -\ c\ {\isaliteral{3D}{\isacharequal}}\ a\ b% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Referencing facts% -} -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Via explicit name:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ a% -\begin{isamarkuptxt}% -Via implicit name:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ this% -\begin{isamarkuptxt}% -Via literal proposition (unification with results from the proof text):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B\ a{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B\ b{\isaliteral{60}{\isacharbackquoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Manipulating facts% -} -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Instantiation:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ a\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ a\ {\isaliteral{5B}{\isacharbrackleft}}of\ b{\isaliteral{5D}{\isacharbrackright}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ a\ {\isaliteral{5B}{\isacharbrackleft}}\isakeyword{where}\ x\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{5D}{\isacharbrackright}}% -\begin{isamarkuptxt}% -Backchaining:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ A\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isadigit{2}}\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isadigit{1}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isadigit{1}}\ {\isaliteral{5B}{\isacharbrackleft}}THEN\ {\isadigit{2}}{\isaliteral{5D}{\isacharbrackright}}% -\begin{isamarkuptxt}% -Symmetric results:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ this\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ this\ {\isaliteral{5B}{\isacharbrackleft}}symmetric{\isaliteral{5D}{\isacharbrackright}}% -\begin{isamarkuptxt}% -Adhoc-simplification (take care!):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5B}{\isacharbrackleft}}{\isaliteral{5D}{\isacharbrackright}}\ {\isaliteral{40}{\isacharat}}\ xs{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ this\ {\isaliteral{5B}{\isacharbrackleft}}simplified{\isaliteral{5D}{\isacharbrackright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Projections% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isar facts consist of multiple theorems. There is notation to project - interval ranges.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ stuff{\isaliteral{3A}{\isacharcolon}}\ A\ B\ C\ D\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{2D}{\isacharminus}}{\isadigit{3}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ stuff{\isaliteral{28}{\isacharparenleft}}{\isadigit{2}}{\isaliteral{2D}{\isacharminus}}{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Naming conventions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{itemize} - - \item Lower-case identifiers are usually preferred. - - \item Facts can be named after the main term within the proposition. - - \item Facts should \emph{not} be named after the command that - introduced them (\hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}, \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}). This is - misleading and hard to maintain. - - \item Natural numbers can be used as ``meaningless'' names (more - appropriate than \isa{{\isaliteral{22}{\isachardoublequote}}a{\isadigit{1}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}a{\isadigit{2}}{\isaliteral{22}{\isachardoublequote}}} etc.) - - \item Symbolic identifiers are supported (e.g. \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}}). - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Block structure% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The formal notepad is block structured. The fact produced by the last - entry of a block is exported into the outer context.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ b{\isaliteral{3A}{\isacharcolon}}\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{note}\isamarkupfalse% -\ a\ b\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ this\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B{\isaliteral{60}{\isacharbackquoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Explicit blocks as well as implicit blocks of nested goal - statements (e.g.\ \hyperlink{command.have}{\mbox{\isa{\isacommand{have}}}}) automatically introduce one extra - pair of parentheses in reserve. The \hyperlink{command.next}{\mbox{\isa{\isacommand{next}}}} command allows - to ``jump'' between these sub-blocks.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\isanewline -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ b{\isaliteral{3A}{\isacharcolon}}\ B\isanewline -\ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ c{\isaliteral{3A}{\isacharcolon}}\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ d{\isaliteral{3A}{\isacharcolon}}\ D\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Alternative version with explicit parentheses everywhere:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ b{\isaliteral{3A}{\isacharcolon}}\ B\isanewline -\ \ \ \ \ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ c{\isaliteral{3A}{\isacharcolon}}\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \ \ \isacommand{have}\isamarkupfalse% -\ d{\isaliteral{3A}{\isacharcolon}}\ D\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsection{Calculational reasoning \label{sec:calculations-synopsis}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -For example, see \verb|~~/src/HOL/Isar_Examples/Group.thy|.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Special names in Isar proofs% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{itemize} - - \item term \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}thesis{\isaliteral{22}{\isachardoublequote}}} --- the main conclusion of the - innermost pending claim - - \item term \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} --- the argument of the last explicitly - stated result (for infix application this is the right-hand side) - - \item fact \isa{{\isaliteral{22}{\isachardoublequote}}this{\isaliteral{22}{\isachardoublequote}}} --- the last result produced in the text - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\ \ \ \ \isacommand{term}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\isanewline -% -\isadelimproof -\ \ \ \ % -\endisadelimproof -% -\isatagproof -\isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\ \ \ \ \isacommand{term}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \ % -\isamarkupcmt{static!% -} -\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\ \ \isacommand{term}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ this\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Calculational reasoning maintains the special fact called - ``\isa{calculation}'' in the background. Certain language - elements combine primary \isa{this} with secondary \isa{calculation}.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Transitive chains% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Idea is to combine \isa{this} and \isa{calculation} - via typical \isa{trans} rules (see also \hyperlink{command.print-trans-rules}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}trans{\isaliteral{5F}{\isacharunderscore}}rules}}}}):% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ trans\isanewline -\isacommand{thm}\isamarkupfalse% -\ less{\isaliteral{5F}{\isacharunderscore}}trans\isanewline -\isacommand{thm}\isamarkupfalse% -\ less{\isaliteral{5F}{\isacharunderscore}}le{\isaliteral{5F}{\isacharunderscore}}trans\isanewline -\isanewline -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -Plain bottom-up calculation:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}b\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}c\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Variant using the \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} abbreviation:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Top-down version with explicit claim at the head:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3D}{\isacharequal}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{finally}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isacommand{next}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Mixed inequalities (require suitable base type):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{fix}\isamarkupfalse% -\ a\ b\ c\ d\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3C}{\isacharless}}\ b{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}b\ {\isaliteral{5C3C6C653E}{\isasymle}}\ c{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}c\ {\isaliteral{3D}{\isacharequal}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{3C}{\isacharless}}\ d{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Notes% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{itemize} - - \item The notion of \isa{trans} rule is very general due to the - flexibility of Isabelle/Pure rule composition. - - \item User applications may declare their own rules, with some care - about the operational details of higher-order unification. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Degenerate calculations and bigstep reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Idea is to append \isa{this} to \isa{calculation}, - without rule composition.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -% -\begin{isamarkuptxt}% -A vacuous proof:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{ultimately}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isakeyword{and}\ C\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{next}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Slightly more content (trivial bigstep reasoning):% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{ultimately}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B\ {\isaliteral{5C3C616E643E}{\isasymand}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ blast\isanewline -\isacommand{next}\isamarkupfalse% -% -\begin{isamarkuptxt}% -More ambitious bigstep reasoning involving structured results:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B\ {\isaliteral{5C3C6F723E}{\isasymor}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ A\ \isacommand{have}\isamarkupfalse% -\ R\ \isacommand{sorry}\isamarkupfalse% -\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ B\ \isacommand{have}\isamarkupfalse% -\ R\ \isacommand{sorry}\isamarkupfalse% -\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{moreover}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\ \isacommand{assume}\isamarkupfalse% -\ C\ \isacommand{have}\isamarkupfalse% -\ R\ \isacommand{sorry}\isamarkupfalse% -\ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{ultimately}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ R\ \isacommand{by}\isamarkupfalse% -\ blast\ \ % -\isamarkupcmt{``big-bang integration'' of proof blocks (occasionally fragile)% -} -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsection{Induction% -} -\isamarkuptrue% -% -\isamarkupsubsection{Induction as Natural Deduction% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -In principle, induction is just a special case of Natural - Deduction (see also \secref{sec:natural-deduction-synopsis}). For - example:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline -\isanewline -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ nat{\isaliteral{2E}{\isachardot}}induct{\isaliteral{29}{\isacharparenright}}\ \ % -\isamarkupcmt{fragile rule application!% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -In practice, much more proof infrastructure is required. - - The proof method \hyperlink{method.induct}{\mbox{\isa{induct}}} provides: - \begin{itemize} - - \item implicit rule selection and robust instantiation - - \item context elements via symbolic case names - - \item support for rule-structured induction statements, with local - parameters, premises, etc. - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ Suc{\isaliteral{2E}{\isachardot}}hyps\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Example% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The subsequent example combines the following proof patterns: - \begin{itemize} - - \item outermost induction (over the datatype structure of natural - numbers), to decompose the proof problem in top-down manner - - \item calculational reasoning (\secref{sec:calculations-synopsis}) - to compose the result in each case - - \item solving local claims within the calculation by simplification - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\isanewline -\ \ \isakeyword{fixes}\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \isakeyword{shows}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -% -\endisadelimproof -% -\isatagproof -\isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}{\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}nat{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isadigit{0}}\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{0}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}Suc\ n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C53756D3E}{\isasymSum}}i{\isaliteral{3D}{\isacharequal}}{\isadigit{0}}{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}n{\isaliteral{2E}{\isachardot}}\ i{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}\ {\isaliteral{2B}{\isacharplus}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}simp\ add{\isaliteral{3A}{\isacharcolon}}\ Suc{\isaliteral{2E}{\isachardot}}hyps{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{also}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C646F74733E}{\isasymdots}}\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ n\ {\isaliteral{2A}{\isacharasterisk}}\ {\isaliteral{28}{\isacharparenleft}}Suc\ n\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{29}{\isacharparenright}}\ div\ {\isadigit{2}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ simp\isanewline -\ \ \isacommand{finally}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}case\ \isacommand{{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\begin{isamarkuptext}% -This demonstrates how induction proofs can be done without - having to consider the raw Natural Deduction structure.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Induction with local parameters and premises% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Idea: Pure rule statements are passed through the induction - rule. This achieves convenient proof patterns, thanks to some - internal trickery in the \hyperlink{method.induct}{\mbox{\isa{induct}}} method. - - Important: Using compact HOL formulae with \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{2F}{\isacharslash}}{\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} is a - well-known anti-pattern! It would produce useless formal noise.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ P\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ Q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}P\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isadigit{0}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ n{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ Q\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ Q\ x\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Local quantification admits arbitrary instances:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}Q\ a\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}Q\ b\ n{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Implicit induction context% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The \hyperlink{method.induct}{\mbox{\isa{induct}}} method can isolate local parameters and - premises directly from the given statement. This is convenient in - practical applications, but requires some understanding of what is - going on internally (as explained above).% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ n\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ nat\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ Q\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ n{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ n\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A\ x\ {\isadigit{0}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isadigit{0}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ n\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ x\ n{\isaliteral{60}{\isacharbackquoteclose}}\ \ % -\isamarkupcmt{arbitrary instances can be produced here% -} -\isanewline -\ \ \ \ \ \ \isakeyword{and}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ x\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Advanced induction with term definitions% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Induction over subexpressions of a certain shape are delicate - to formalize. The Isar \hyperlink{method.induct}{\mbox{\isa{induct}}} method provides - infrastructure for this. - - Idea: sub-expressions of the problem are turned into a defined - induction variable; often accompanied with fixing of auxiliary - parameters in the original expression.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ a\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ nat{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{fix}\isamarkupfalse% -\ A\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}nat\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}induct\ {\isaliteral{22}{\isachardoublequoteopen}}a\ x{\isaliteral{22}{\isachardoublequoteclose}}\ arbitrary{\isaliteral{3A}{\isacharcolon}}\ x{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isadigit{0}}\isanewline -\ \ \ \ \isacommand{note}\isamarkupfalse% -\ prem\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ defn\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}{\isadigit{0}}\ {\isaliteral{3D}{\isacharequal}}\ a\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Suc\ n{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{note}\isamarkupfalse% -\ hyp\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ n\ {\isaliteral{3D}{\isacharequal}}\ a\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ prem\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}A\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \ \ \isakeyword{and}\ defn\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{60}{\isacharbackquoteopen}}Suc\ n\ {\isaliteral{3D}{\isacharequal}}\ a\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{28}{\isacharparenleft}}a\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsection{Natural Deduction \label{sec:natural-deduction-synopsis}% -} -\isamarkuptrue% -% -\isamarkupsubsection{Rule statements% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/Pure ``theorems'' are always natural deduction rules, - which sometimes happen to consist of a conclusion only. - - The framework connectives \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}{\isaliteral{22}{\isachardoublequote}}} and \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}{\isaliteral{22}{\isachardoublequote}}} indicate the - rule structure declaratively. For example:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ conjI\isanewline -\isacommand{thm}\isamarkupfalse% -\ impI\isanewline -\isacommand{thm}\isamarkupfalse% -\ nat{\isaliteral{2E}{\isachardot}}induct% -\begin{isamarkuptext}% -The object-logic is embedded into the Pure framework via an implicit - derivability judgment \isa{{\isaliteral{22}{\isachardoublequote}}Trueprop\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ bool\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop{\isaliteral{22}{\isachardoublequote}}}. - - Thus any HOL formulae appears atomic to the Pure framework, while - the rule structure outlines the corresponding proof pattern. - - This can be made explicit as follows:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{write}\isamarkupfalse% -\ Trueprop\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{22}{\isachardoublequoteopen}}Tr{\isaliteral{22}{\isachardoublequoteclose}}{\isaliteral{29}{\isacharparenright}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ impI\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ nat{\isaliteral{2E}{\isachardot}}induct\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Isar provides first-class notation for rule statements as follows.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ conjI\isanewline -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ impI\isanewline -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ nat{\isaliteral{2E}{\isachardot}}induct% -\isamarkupsubsubsection{Examples% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Introductions and eliminations of some standard connectives of - the object-logic can be written as rule statements as follows. (The - proof ``\hyperlink{command.by}{\mbox{\isa{\isacommand{by}}}}~\hyperlink{method.blast}{\mbox{\isa{blast}}}'' serves as sanity check.)% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ False{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C616E643E}{\isasymand}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ {\isaliteral{5C3C6F723E}{\isasymor}}\ Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}P\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}Q\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ P\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isanewline -\isacommand{lemma}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{22}{\isachardoublequoteclose}}% -\isadelimproof -\ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsection{Isar context elements% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -We derive some results out of the blue, using Isar context - elements and some explicit blocks. This illustrates their meaning - wrt.\ Pure connectives, without goal states getting in the way.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{def}\isamarkupfalse% -\ x\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ t\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ t{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{by}\isamarkupfalse% -\ fact\isanewline -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{obtain}\isamarkupfalse% -\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{by}\isamarkupfalse% -\ fact% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Pure rule composition% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The Pure framework provides means for: - - \begin{itemize} - - \item backward-chaining of rules by \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} - - \item closing of branches by \hyperlink{inference.assumption}{\mbox{\isa{assumption}}} - - \end{itemize} - - Both principles involve higher-order unification of \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-terms - modulo \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616C7068613E}{\isasymalpha}}{\isaliteral{5C3C626574613E}{\isasymbeta}}{\isaliteral{5C3C6574613E}{\isasymeta}}{\isaliteral{22}{\isachardoublequote}}}-equivalence (cf.\ Huet and Miller).% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ a{\isaliteral{3A}{\isacharcolon}}\ A\ \isakeyword{and}\ b{\isaliteral{3A}{\isacharcolon}}\ B% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}of\ A\ B{\isaliteral{5D}{\isacharbrackright}}\ \ % -\isamarkupcmt{instantiation% -} -\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}of\ A\ B{\isaliteral{2C}{\isacharcomma}}\ OF\ a\ b{\isaliteral{5D}{\isacharbrackright}}\ \ % -\isamarkupcmt{instantiation and composition% -} -\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ a\ b{\isaliteral{5D}{\isacharbrackright}}\ \ % -\isamarkupcmt{composition via unification (trivial)% -} -\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\ {\isaliteral{60}{\isacharbackquoteopen}}B{\isaliteral{60}{\isacharbackquoteclose}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\isanewline -\ \ \isacommand{thm}\isamarkupfalse% -\ conjI\ {\isaliteral{5B}{\isacharbrackleft}}OF\ disjI{\isadigit{1}}{\isaliteral{5D}{\isacharbrackright}}\isanewline -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Note: Low-level rule composition is tedious and leads to - unreadable~/ unmaintainable expressions in the text.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Structured backward reasoning% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Idea: Canonical proof decomposition via \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~/ - \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~/ \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}, where the body produces a - natural deduction rule to refine some goal.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ A\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{2D}{\isacharminus}}\isanewline -\ \ \ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\ % -\isamarkupcmt{implicit block structure made explicit% -} -\isanewline -\ \ \ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ A\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ x{\isaliteral{60}{\isacharbackquoteclose}}\isanewline -\ \ \ \ \ \ % -\isamarkupcmt{side exit for the resulting rule% -} -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Structured rule application% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Idea: Previous facts and new claims are composed with a rule from - the context (or background library).% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ r{\isadigit{1}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \ % -\isamarkupcmt{simple rule (Horn clause)% -} -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\ \ % -\isamarkupcmt{prefix of facts via outer sub-proof% -} -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\ \ % -\isamarkupcmt{remaining rule premises via inner sub-proof% -} -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{by}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\isanewline -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{assume}\isamarkupfalse% -\ r{\isadigit{2}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isadigit{2}}\ x{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isaliteral{22}{\isachardoublequoteclose}}\ \ % -\isamarkupcmt{nested rule% -} -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\begin{isamarkuptxt}% -The compound rule premise \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B{\isadigit{2}}\ x{\isaliteral{22}{\isachardoublequote}}} is better - addressed via \hyperlink{command.fix}{\mbox{\isa{\isacommand{fix}}}}~/ \hyperlink{command.assume}{\mbox{\isa{\isacommand{assume}}}}~/ \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}} - in the nested proof body.% -\end{isamarkuptxt}% -\isamarkuptrue% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Example: predicate logic% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Using the above principles, standard introduction and elimination proofs - of predicate logic connectives of HOL work as follows.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6C6F6E6772696768746172726F773E}{\isasymlongrightarrow}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C6F723E}{\isasymor}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ B\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ False\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ A\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ True\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ A\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ False\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ B\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ a\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\begin{isamarkuptxt}% -Less awkward version using \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}:% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}x{\isaliteral{2E}{\isachardot}}\ P\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}P\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Further variations to illustrate Isar sub-proofs involving - \hyperlink{command.show}{\mbox{\isa{\isacommand{show}}}}:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \ % -\isamarkupcmt{two strictly isolated subproofs% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \ % -\isamarkupcmt{one simultaneous sub-proof% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isakeyword{and}\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \ % -\isamarkupcmt{two subproofs in the same context% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \ % -\isamarkupcmt{swapped order% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}A\ {\isaliteral{5C3C616E643E}{\isasymand}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ \ % -\isamarkupcmt{sequential subproofs% -} -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ A\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ B\ \isacommand{using}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}A{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Example: set-theoretic operators% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -There is nothing special about logical connectives (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C616E643E}{\isasymand}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6F723E}{\isasymor}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C666F72616C6C3E}{\isasymforall}}{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6578697374733E}{\isasymexists}}{\isaliteral{22}{\isachardoublequote}}} etc.). Operators from - set-theory or lattice-theory work analogously. It is only a matter - of rule declarations in the library; rules can be also specified - explicitly.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C696E7465723E}{\isasyminter}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A\ {\isaliteral{5C3C756E696F6E3E}{\isasymunion}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ B{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{show}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\isacommand{next}\isamarkupfalse% -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ a\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C496E7465723E}{\isasymInter}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ {\isaliteral{5C3C556E696F6E3E}{\isasymUnion}}A{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{obtain}\isamarkupfalse% -\ a\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}a\ {\isaliteral{5C3C696E3E}{\isasymin}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{5C3C696E3E}{\isasymin}}\ a{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{{\isaliteral{2E}{\isachardot}}{\isaliteral{2E}{\isachardot}}}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsection{Generalized elimination and cases% -} -\isamarkuptrue% -% -\isamarkupsubsection{General elimination rules% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -The general format of elimination rules is illustrated by the - following typical representatives:% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{thm}\isamarkupfalse% -\ exE\ \ \ \ \ % -\isamarkupcmt{local parameter% -} -\isanewline -\isacommand{thm}\isamarkupfalse% -\ conjE\ \ \ % -\isamarkupcmt{local premises% -} -\isanewline -\isacommand{thm}\isamarkupfalse% -\ disjE\ \ \ % -\isamarkupcmt{split into cases% -} -% -\begin{isamarkuptext}% -Combining these characteristics leads to the following general scheme - for elimination rules with cases: - - \begin{itemize} - - \item prefix of assumptions (or ``major premises'') - - \item one or more cases that enable to establish the main conclusion - in an augmented context - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{assume}\isamarkupfalse% -\ r{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \ \ {\isaliteral{22}{\isachardoublequoteopen}}A{\isadigit{1}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ A{\isadigit{2}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ assumptions\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ B{\isadigit{1}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isadigit{1}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ case\ {\isadigit{1}}\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x\ y{\isaliteral{2E}{\isachardot}}\ B{\isadigit{2}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ C{\isadigit{2}}\ x\ y\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ R{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ case\ {\isadigit{2}}\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \ \ R\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}\ main\ conclusion\ {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ A{\isadigit{1}}\ \isakeyword{and}\ A{\isadigit{2}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{then}\isamarkupfalse% -\ \isacommand{have}\isamarkupfalse% -\ R\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}rule\ r{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{fix}\isamarkupfalse% -\ x\ y\isanewline -\ \ \ \ \isacommand{assume}\isamarkupfalse% -\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ \ \ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\begin{isamarkuptext}% -Here \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{3F}{\isacharquery}}thesis{\isaliteral{22}{\isachardoublequote}}} is used to refer to the unchanged goal - statement.% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsubsection{Rules with cases% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Applying an elimination rule to some goal, leaves that unchanged - but allows to augment the context in the sub-proof of each case. - - Isar provides some infrastructure to support this: - - \begin{itemize} - - \item native language elements to state eliminations - - \item symbolic case names - - \item method \hyperlink{method.cases}{\mbox{\isa{cases}}} to recover this structure in a - sub-proof - - \end{itemize}% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ exE\isanewline -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ conjE\isanewline -\isacommand{print{\isaliteral{5F}{\isacharunderscore}}statement}\isamarkupfalse% -\ disjE\isanewline -\isanewline -\isacommand{lemma}\isamarkupfalse% -\isanewline -\ \ \isakeyword{assumes}\ A{\isadigit{1}}\ \isakeyword{and}\ A{\isadigit{2}}\ \ % -\isamarkupcmt{assumptions% -} -\isanewline -\ \ \isakeyword{obtains}\isanewline -\ \ \ \ {\isaliteral{28}{\isacharparenleft}}case{\isadigit{1}}{\isaliteral{29}{\isacharparenright}}\ \ x\ y\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{1}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{28}{\isacharparenleft}}case{\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ \ x\ y\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\ \isakeyword{and}\ {\isaliteral{22}{\isachardoublequoteopen}}C{\isadigit{2}}\ x\ y{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{sorry}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -% -\endisadelimproof -% -\isamarkupsubsubsection{Example% -} -\isamarkuptrue% -\isacommand{lemma}\isamarkupfalse% -\ tertium{\isaliteral{5F}{\isacharunderscore}}non{\isaliteral{5F}{\isacharunderscore}}datur{\isaliteral{3A}{\isacharcolon}}\isanewline -\ \ \isakeyword{obtains}\isanewline -\ \ \ \ {\isaliteral{28}{\isacharparenleft}}T{\isaliteral{29}{\isacharparenright}}\ \ A\isanewline -\ \ {\isaliteral{7C}{\isacharbar}}\ {\isaliteral{28}{\isacharparenleft}}F{\isaliteral{29}{\isacharparenright}}\ \ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{5C3C6E6F743E}{\isasymnot}}\ A{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{by}\isamarkupfalse% -\ blast% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isanewline -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ x\ y\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{27}{\isacharprime}}a\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}cases\ {\isaliteral{22}{\isachardoublequoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{22}{\isachardoublequoteclose}}\ rule{\isaliteral{3A}{\isacharcolon}}\ tertium{\isaliteral{5F}{\isacharunderscore}}non{\isaliteral{5F}{\isacharunderscore}}datur{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ T\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ F\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{5C3C6E6F7465713E}{\isasymnoteq}}\ y{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsubsection{Example% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -Isabelle/HOL specification mechanisms (datatype, inductive, etc.) - provide suitable derived cases rules.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{datatype}\isamarkupfalse% -\ foo\ {\isaliteral{3D}{\isacharequal}}\ Foo\ {\isaliteral{7C}{\isacharbar}}\ Bar\ foo\isanewline -\isanewline -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ x\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ foo\isanewline -\ \ \isacommand{have}\isamarkupfalse% -\ C\isanewline -\ \ \isacommand{proof}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}cases\ x{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ Foo\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Foo{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{next}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{case}\isamarkupfalse% -\ {\isaliteral{28}{\isacharparenleft}}Bar\ a{\isaliteral{29}{\isacharparenright}}\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}x\ {\isaliteral{3D}{\isacharequal}}\ Bar\ a{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{show}\isamarkupfalse% -\ {\isaliteral{3F}{\isacharquery}}thesis\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{qed}\isamarkupfalse% -% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -% -\isamarkupsubsection{Obtaining local contexts% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -A single ``case'' branch may be inlined into Isar proof text - via \hyperlink{command.obtain}{\mbox{\isa{\isacommand{obtain}}}}. This proves \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}{\isaliteral{5C3C416E643E}{\isasymAnd}}x{\isaliteral{2E}{\isachardot}}\ B\ x\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{29}{\isacharparenright}}\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ thesis{\isaliteral{22}{\isachardoublequote}}} on the spot, and augments the context afterwards.% -\end{isamarkuptext}% -\isamarkuptrue% -\isacommand{notepad}\isamarkupfalse% -\isanewline -\isakeyword{begin}\isanewline -% -\isadelimproof -\ \ % -\endisadelimproof -% -\isatagproof -\isacommand{fix}\isamarkupfalse% -\ B\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{22}{\isachardoublequoteopen}}{\isaliteral{27}{\isacharprime}}a\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ bool{\isaliteral{22}{\isachardoublequoteclose}}\isanewline -\isanewline -\ \ \isacommand{obtain}\isamarkupfalse% -\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}% -\begin{isamarkuptxt}% -Conclusions from this context may not mention \isa{x} again!% -\end{isamarkuptxt}% -\isamarkuptrue% -\ \ \isacommand{{\isaliteral{7B}{\isacharbraceleft}}}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{obtain}\isamarkupfalse% -\ x\ \isakeyword{where}\ {\isaliteral{22}{\isachardoublequoteopen}}B\ x{\isaliteral{22}{\isachardoublequoteclose}}\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \ \ \isacommand{from}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}B\ x{\isaliteral{60}{\isacharbackquoteclose}}\ \isacommand{have}\isamarkupfalse% -\ C\ \isacommand{sorry}\isamarkupfalse% -\isanewline -\ \ \isacommand{{\isaliteral{7D}{\isacharbraceright}}}\isamarkupfalse% -\isanewline -\ \ \isacommand{note}\isamarkupfalse% -\ {\isaliteral{60}{\isacharbackquoteopen}}C{\isaliteral{60}{\isacharbackquoteclose}}% -\endisatagproof -{\isafoldproof}% -% -\isadelimproof -\isanewline -% -\endisadelimproof -\isacommand{end}\isamarkupfalse% -\isanewline -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -% -\isatagtheory -\isacommand{end}\isamarkupfalse% -% -\endisatagtheory -{\isafoldtheory}% -% -\isadelimtheory -\isanewline -% -\endisadelimtheory -\end{isabellebody}% -%%% Local Variables: -%%% mode: latex -%%% TeX-master: "root" -%%% End: diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/isar-vm.eps --- a/doc-src/IsarRef/Thy/document/isar-vm.eps Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2694 +0,0 @@ -%!PS-Adobe-3.0 EPSF-3.0 -%%Creator: inkscape 0.46 -%%Pages: 1 -%%Orientation: Portrait -%%BoundingBox: 0 0 435 173 -%%HiResBoundingBox: 0 0 435 173 -%%EndComments -%%BeginSetup -%%EndSetup -%%Page: 1 1 -0 173 translate -0.8 -0.8 scale -0 0 0 setrgbcolor -[] 0 setdash -1 setlinewidth -0 setlinejoin -0 setlinecap -gsave [1 0 0 1 0 0] concat -gsave [1 0 0 1 -44.641342 -76.87234] concat -gsave [1 0 0 1 70.838012 79.725562] concat -0 0 0 setrgbcolor -[] 0 setdash -0.99921262 setlinewidth -1 setlinejoin -1 setlinecap -newpath -229.77649 131.52507 moveto -265.28729 131.52507 lineto -275.08072 131.52507 282.96496 139.40931 282.96496 149.20274 curveto -282.96496 166.99701 lineto -282.96496 176.79043 275.08072 184.67467 265.28729 184.67467 curveto -229.77649 184.67467 lineto -219.98306 184.67467 212.09882 176.79043 212.09882 166.99701 curveto -212.09882 149.20274 lineto -212.09882 139.40931 219.98306 131.52507 229.77649 131.52507 curveto -closepath -stroke -gsave -0 0 0 setrgbcolor -newpath -231.92252 155.58815 moveto -231.92252 157.8694 lineto -231.5423 157.60899 231.15949 157.41628 230.77408 157.29128 curveto -230.39386 157.16628 229.99803 157.10378 229.58658 157.10378 curveto -228.80532 157.10378 228.19595 157.33295 227.75845 157.79128 curveto -227.32616 158.24441 227.11001 158.87982 227.11002 159.69753 curveto -227.11001 160.51524 227.32616 161.15326 227.75845 161.61159 curveto -228.19595 162.06471 228.80532 162.29128 229.58658 162.29128 curveto -230.02407 162.29128 230.43813 162.22617 230.82877 162.09596 curveto -231.22459 161.96576 231.58917 161.77305 231.92252 161.51784 curveto -231.92252 163.8069 lineto -231.48501 163.96836 231.0397 164.08815 230.58658 164.16628 curveto -230.13866 164.24961 229.68813 164.29127 229.23502 164.29128 curveto -227.65689 164.29127 226.42251 163.88763 225.53189 163.08034 curveto -224.64126 162.26784 224.19595 161.14024 224.19595 159.69753 curveto -224.19595 158.25482 224.64126 157.12982 225.53189 156.32253 curveto -226.42251 155.51003 227.65689 155.10378 229.23502 155.10378 curveto -229.69334 155.10378 230.14386 155.14545 230.58658 155.22878 curveto -231.03449 155.30691 231.4798 155.4267 231.92252 155.58815 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -243.14908 158.73659 moveto -243.14908 164.06471 lineto -240.33658 164.06471 lineto -240.33658 163.19753 lineto -240.33658 160.00221 lineto -240.33657 159.23659 240.31834 158.71055 240.28189 158.42409 curveto -240.25063 158.13764 240.19334 157.9267 240.11002 157.79128 curveto -240.00063 157.60899 239.8522 157.46836 239.6647 157.3694 curveto -239.4772 157.26524 239.26366 157.21316 239.02408 157.21315 curveto -238.44074 157.21316 237.98241 157.43972 237.64908 157.89284 curveto -237.31574 158.34076 237.14907 158.96316 237.14908 159.76003 curveto -237.14908 164.06471 lineto -234.3522 164.06471 lineto -234.3522 151.90846 lineto -237.14908 151.90846 lineto -237.14908 156.59596 lineto -237.57095 156.08555 238.01887 155.71055 238.49283 155.47096 curveto -238.96678 155.22618 239.49022 155.10378 240.06314 155.10378 curveto -241.07355 155.10378 241.83917 155.41368 242.36002 156.03346 curveto -242.88605 156.65326 243.14907 157.5543 243.14908 158.73659 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -249.68033 160.12721 moveto -249.09699 160.12722 248.65689 160.22617 248.36002 160.42409 curveto -248.06835 160.62201 247.92251 160.91367 247.92252 161.29909 curveto -247.92251 161.65326 248.0397 161.9319 248.27408 162.13503 curveto -248.51366 162.33294 248.84439 162.4319 249.26627 162.4319 curveto -249.7923 162.4319 250.23501 162.2444 250.59439 161.8694 curveto -250.95376 161.48919 251.13345 161.01524 251.13345 160.44753 curveto -251.13345 160.12721 lineto -249.68033 160.12721 lineto -253.95377 159.07253 moveto -253.95377 164.06471 lineto -251.13345 164.06471 lineto -251.13345 162.76784 lineto -250.75845 163.29909 250.33657 163.68711 249.86783 163.9319 curveto -249.39907 164.17148 248.82876 164.29127 248.15689 164.29128 curveto -247.25064 164.29127 246.51366 164.02825 245.94595 163.50221 curveto -245.38345 162.97096 245.1022 162.28346 245.1022 161.43971 curveto -245.1022 160.41367 245.45376 159.66107 246.15689 159.1819 curveto -246.86522 158.70274 247.9746 158.46316 249.48502 158.46315 curveto -251.13345 158.46315 lineto -251.13345 158.2444 lineto -251.13345 157.8017 250.95897 157.47878 250.61002 157.27565 curveto -250.26105 157.06732 249.71678 156.96316 248.9772 156.96315 curveto -248.37824 156.96316 247.82095 157.02305 247.30533 157.14284 curveto -246.7897 157.26264 246.31053 157.44232 245.86783 157.6819 curveto -245.86783 155.54909 lineto -246.46678 155.40326 247.06835 155.29389 247.67252 155.22096 curveto -248.27668 155.14285 248.88084 155.10378 249.48502 155.10378 curveto -251.06313 155.10378 252.20115 155.41628 252.89908 156.04128 curveto -253.60219 156.66107 253.95376 157.67149 253.95377 159.07253 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -256.57095 155.31471 moveto -259.36783 155.31471 lineto -259.36783 164.06471 lineto -256.57095 164.06471 lineto -256.57095 155.31471 lineto -256.57095 151.90846 moveto -259.36783 151.90846 lineto -259.36783 154.18971 lineto -256.57095 154.18971 lineto -256.57095 151.90846 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -270.86783 158.73659 moveto -270.86783 164.06471 lineto -268.05533 164.06471 lineto -268.05533 163.19753 lineto -268.05533 159.98659 lineto -268.05532 159.23138 268.03709 158.71055 268.00064 158.42409 curveto -267.96938 158.13764 267.91209 157.9267 267.82877 157.79128 curveto -267.71938 157.60899 267.57095 157.46836 267.38345 157.3694 curveto -267.19595 157.26524 266.98241 157.21316 266.74283 157.21315 curveto -266.15949 157.21316 265.70116 157.43972 265.36783 157.89284 curveto -265.03449 158.34076 264.86782 158.96316 264.86783 159.76003 curveto -264.86783 164.06471 lineto -262.07095 164.06471 lineto -262.07095 155.31471 lineto -264.86783 155.31471 lineto -264.86783 156.59596 lineto -265.2897 156.08555 265.73762 155.71055 266.21158 155.47096 curveto -266.68553 155.22618 267.20897 155.10378 267.78189 155.10378 curveto -268.7923 155.10378 269.55792 155.41368 270.07877 156.03346 curveto -270.6048 156.65326 270.86782 157.5543 270.86783 158.73659 curveto -fill -grestore -grestore -0 0 0 setrgbcolor -[] 0 setdash -0.99921262 setlinewidth -0 setlinejoin -0 setlinecap -newpath -424.72469 236.82544 moveto -356.83209 236.82544 lineto -356.83209 236.82544 lineto -stroke -gsave [-0.39968505 4.8945685e-17 -4.8945685e-17 -0.39968505 356.83209 236.82544] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -0 0 0 setrgbcolor -[] 0 setdash -0.99921268 setlinewidth -0 setlinejoin -0 setlinecap -newpath -282.35183 236.82544 moveto -215.11403 236.82544 lineto -215.11403 236.82544 lineto -stroke -gsave [-0.39968507 4.8945688e-17 -4.8945688e-17 -0.39968507 215.11403 236.82544] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -0 0 0 setrgbcolor -[] 0 setdash -0.99999994 setlinewidth -0 setlinejoin -0 setlinecap -newpath -424.69726 192.5341 moveto -215.13005 192.5341 lineto -stroke -gsave [-0.39999998 4.8984251e-17 -4.8984251e-17 -0.39999998 215.13005 192.5341] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -0 0 0 setrgbcolor -[] 0 setdash -1 setlinewidth -0 setlinejoin -0 setlinecap -newpath -211.98429 148.24276 moveto -422.13162 148.24276 lineto -stroke -gsave [0.4 0 0 0.4 422.13162 148.24276] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -gsave [1 0 0 1 70.866146 78.725567] concat -0 0 0 setrgbcolor -[] 0 setdash -0.99921262 setlinewidth -1 setlinejoin -1 setlinecap -newpath -88.044201 42.942394 moveto -123.555 42.942394 lineto -133.34843 42.942394 141.23267 50.826635 141.23267 60.620064 curveto -141.23267 166.99701 lineto -141.23267 176.79044 133.34843 184.67468 123.555 184.67468 curveto -88.044201 184.67468 lineto -78.250772 184.67468 70.366531 176.79044 70.366531 166.99701 curveto -70.366531 60.620064 lineto -70.366531 50.826635 78.250772 42.942394 88.044201 42.942394 curveto -closepath -stroke -gsave -0 0 0 setrgbcolor -newpath -83.823044 115.35931 moveto -83.823044 119.95306 lineto -81.026169 119.95306 lineto -81.026169 107.87494 lineto -83.823044 107.87494 lineto -83.823044 109.15619 lineto -84.208456 108.64578 84.635539 108.27078 85.104294 108.03119 curveto -85.573038 107.78641 86.1121 107.66401 86.721481 107.664 curveto -87.799598 107.66401 88.685014 108.0937 89.377731 108.95306 curveto -90.070429 109.80724 90.416783 110.9088 90.416794 112.25775 curveto -90.416783 113.60671 90.070429 114.71088 89.377731 115.57025 curveto -88.685014 116.42442 87.799598 116.8515 86.721481 116.8515 curveto -86.1121 116.8515 85.573038 116.73171 85.104294 116.49213 curveto -84.635539 116.24734 84.208456 115.86973 83.823044 115.35931 curveto -85.682419 109.69525 moveto -85.083455 109.69526 84.622518 109.91661 84.299606 110.35931 curveto -83.981894 110.79682 83.82304 111.42963 83.823044 112.25775 curveto -83.82304 113.08588 83.981894 113.7213 84.299606 114.164 curveto -84.622518 114.6015 85.083455 114.82025 85.682419 114.82025 curveto -86.281371 114.82025 86.737099 114.6015 87.049606 114.164 curveto -87.367307 113.7265 87.526161 113.09109 87.526169 112.25775 curveto -87.526161 111.42442 87.367307 110.78901 87.049606 110.3515 curveto -86.737099 109.91401 86.281371 109.69526 85.682419 109.69525 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -98.994919 110.25775 moveto -98.75012 110.14317 98.505328 110.05984 98.260544 110.00775 curveto -98.020954 109.95047 97.778766 109.92182 97.533981 109.92181 curveto -96.815226 109.92182 96.260539 110.15359 95.869919 110.61713 curveto -95.484498 111.07547 95.29179 111.73432 95.291794 112.59369 curveto -95.291794 116.62494 lineto -92.494919 116.62494 lineto -92.494919 107.87494 lineto -95.291794 107.87494 lineto -95.291794 109.31244 lineto -95.651164 108.73953 96.062622 108.32286 96.526169 108.06244 curveto -96.994913 107.79682 97.554808 107.66401 98.205856 107.664 curveto -98.299599 107.66401 98.401162 107.66922 98.510544 107.67963 curveto -98.619911 107.68484 98.778765 107.70047 98.987106 107.7265 curveto -98.994919 110.25775 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -104.56523 109.664 moveto -103.94543 109.66401 103.47148 109.88797 103.14336 110.33588 curveto -102.82044 110.77859 102.65898 111.41922 102.65898 112.25775 curveto -102.65898 113.0963 102.82044 113.73953 103.14336 114.18744 curveto -103.47148 114.63015 103.94543 114.8515 104.56523 114.8515 curveto -105.1746 114.8515 105.64075 114.63015 105.96367 114.18744 curveto -106.28658 113.73953 106.44804 113.0963 106.44804 112.25775 curveto -106.44804 111.41922 106.28658 110.77859 105.96367 110.33588 curveto -105.64075 109.88797 105.1746 109.66401 104.56523 109.664 curveto -104.56523 107.664 moveto -106.07043 107.66401 107.24491 108.07026 108.08867 108.88275 curveto -108.93762 109.69526 109.3621 110.82026 109.36211 112.25775 curveto -109.3621 113.69525 108.93762 114.82025 108.08867 115.63275 curveto -107.24491 116.44525 106.07043 116.8515 104.56523 116.8515 curveto -103.05481 116.8515 101.87252 116.44525 101.01836 115.63275 curveto -100.1694 114.82025 99.744918 113.69525 99.744919 112.25775 curveto -99.744918 110.82026 100.1694 109.69526 101.01836 108.88275 curveto -101.87252 108.07026 103.05481 107.66401 104.56523 107.664 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -110.29961 107.87494 moveto -113.09648 107.87494 lineto -115.27617 113.92181 lineto -117.44804 107.87494 lineto -120.25273 107.87494 lineto -116.80742 116.62494 lineto -113.73711 116.62494 lineto -110.29961 107.87494 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -130.57304 112.2265 moveto -130.57304 113.02338 lineto -124.03398 113.02338 lineto -124.10169 113.67963 124.33866 114.17182 124.74492 114.49994 curveto -125.15116 114.82807 125.71887 114.99213 126.44804 114.99213 curveto -127.03658 114.99213 127.63814 114.90619 128.25273 114.73431 curveto -128.87251 114.55723 129.50793 114.29161 130.15898 113.93744 curveto -130.15898 116.09369 lineto -129.49751 116.34369 128.83606 116.53119 128.17461 116.65619 curveto -127.51314 116.7864 126.85168 116.8515 126.19023 116.8515 curveto -124.60689 116.8515 123.37512 116.45046 122.49492 115.64838 curveto -121.61992 114.84109 121.18242 113.71088 121.18242 112.25775 curveto -121.18242 110.83067 121.61211 109.70828 122.47148 108.89056 curveto -123.33606 108.07286 124.52356 107.66401 126.03398 107.664 curveto -127.40897 107.66401 128.50793 108.07807 129.33086 108.90619 curveto -130.15897 109.73432 130.57303 110.84109 130.57304 112.2265 curveto -127.69804 111.29681 moveto -127.69804 110.76557 127.54179 110.33849 127.22929 110.01556 curveto -126.922 109.68745 126.51835 109.52338 126.01836 109.52338 curveto -125.47668 109.52338 125.03658 109.67703 124.69804 109.98431 curveto -124.3595 110.2864 124.14856 110.7239 124.06523 111.29681 curveto -127.69804 111.29681 lineto -fill -grestore -grestore -0 0 0 setrgbcolor -[] 0 setdash -1 setlinewidth -0 setlinejoin -0 setlinecap -newpath -176.66575 92.035445 moveto -176.66575 118.61025 lineto -stroke -gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 176.66575 118.61025] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -gsave [0.2378166 0 0 -0.2269133 90.621413 253.06251] concat -0 0 0 setrgbcolor -[] 0 setdash -4.3013706 setlinewidth -1 setlinejoin -1 setlinecap -newpath -208.65508 282.05865 moveto -193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto -43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto -45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto -177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto -stroke -gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -grestore -gsave [1 0 0 1 70.866151 78.725565] concat -0 0 0 setrgbcolor -[] 0 setdash -0.99921262 setlinewidth -1 setlinejoin -1 setlinecap -newpath -371.50879 42.942394 moveto -407.01959 42.942394 lineto -416.81302 42.942394 424.69726 50.826635 424.69726 60.620064 curveto -424.69726 166.99701 lineto -424.69726 176.79044 416.81302 184.67468 407.01959 184.67468 curveto -371.50879 184.67468 lineto -361.71536 184.67468 353.83112 176.79044 353.83112 166.99701 curveto -353.83112 60.620064 lineto -353.83112 50.826635 361.71536 42.942394 371.50879 42.942394 curveto -closepath -stroke -gsave -0 0 0 setrgbcolor -newpath -374.16263 110.83588 moveto -374.16263 112.96088 lineto -373.56366 112.71088 372.98554 112.52338 372.42825 112.39838 curveto -371.87096 112.27338 371.34491 112.21088 370.85013 112.21088 curveto -370.31887 112.21088 369.92304 112.27859 369.66263 112.414 curveto -369.40742 112.54422 369.27981 112.74734 369.27982 113.02338 curveto -369.27981 113.24734 369.37617 113.41922 369.56888 113.539 curveto -369.76679 113.6588 370.11835 113.74734 370.62357 113.80463 curveto -371.11575 113.87494 lineto -372.54804 114.05724 373.51158 114.35671 374.00638 114.77338 curveto -374.50116 115.19005 374.74856 115.84369 374.74857 116.73431 curveto -374.74856 117.66661 374.40481 118.36713 373.71732 118.83588 curveto -373.02981 119.30463 372.00377 119.539 370.63919 119.539 curveto -370.06106 119.539 369.4621 119.49213 368.84232 119.39838 curveto -368.22773 119.30983 367.59492 119.17442 366.94388 118.99213 curveto -366.94388 116.86713 lineto -367.50117 117.13796 368.07148 117.34109 368.65482 117.4765 curveto -369.24335 117.61192 369.83971 117.67963 370.44388 117.67963 curveto -370.99075 117.67963 371.40221 117.60411 371.67825 117.45306 curveto -371.95429 117.30202 372.09231 117.07807 372.09232 116.78119 curveto -372.09231 116.53119 371.99596 116.3463 371.80325 116.2265 curveto -371.61575 116.1015 371.23814 116.00515 370.67044 115.93744 curveto -370.17825 115.87494 lineto -368.93346 115.71869 368.06106 115.42963 367.56107 115.00775 curveto -367.06106 114.58588 366.81106 113.94526 366.81107 113.08588 curveto -366.81106 112.1588 367.12877 111.4713 367.76419 111.02338 curveto -368.3996 110.57547 369.37356 110.35151 370.68607 110.3515 curveto -371.20169 110.35151 371.74335 110.39057 372.31107 110.46869 curveto -372.87877 110.54682 373.49595 110.66922 374.16263 110.83588 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -379.91263 108.07806 moveto -379.91263 110.56244 lineto -382.79544 110.56244 lineto -382.79544 112.56244 lineto -379.91263 112.56244 lineto -379.91263 116.27338 lineto -379.91262 116.67963 379.99335 116.95567 380.15482 117.1015 curveto -380.31627 117.24213 380.63658 117.31244 381.11575 117.31244 curveto -382.55325 117.31244 lineto -382.55325 119.31244 lineto -380.15482 119.31244 lineto -379.05065 119.31244 378.26679 119.08327 377.80325 118.62494 curveto -377.34492 118.1614 377.11575 117.37755 377.11575 116.27338 curveto -377.11575 112.56244 lineto -375.72513 112.56244 lineto -375.72513 110.56244 lineto -377.11575 110.56244 lineto -377.11575 108.07806 lineto -379.91263 108.07806 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -388.43607 115.37494 moveto -387.85273 115.37494 387.41262 115.4739 387.11575 115.67181 curveto -386.82408 115.86973 386.67825 116.1614 386.67825 116.54681 curveto -386.67825 116.90098 386.79544 117.17963 387.02982 117.38275 curveto -387.26939 117.58067 387.60012 117.67963 388.022 117.67963 curveto -388.54804 117.67963 388.99075 117.49213 389.35013 117.11713 curveto -389.7095 116.73692 389.88918 116.26296 389.88919 115.69525 curveto -389.88919 115.37494 lineto -388.43607 115.37494 lineto -392.7095 114.32025 moveto -392.7095 119.31244 lineto -389.88919 119.31244 lineto -389.88919 118.01556 lineto -389.51418 118.54681 389.09231 118.93484 388.62357 119.17963 curveto -388.15481 119.41921 387.5845 119.539 386.91263 119.539 curveto -386.00638 119.539 385.2694 119.27598 384.70169 118.74994 curveto -384.13919 118.21869 383.85794 117.53119 383.85794 116.68744 curveto -383.85794 115.6614 384.2095 114.9088 384.91263 114.42963 curveto -385.62096 113.95047 386.73033 113.71088 388.24075 113.71088 curveto -389.88919 113.71088 lineto -389.88919 113.49213 lineto -389.88918 113.04942 389.7147 112.72651 389.36575 112.52338 curveto -389.01679 112.31505 388.47252 112.21088 387.73294 112.21088 curveto -387.13398 112.21088 386.57669 112.27078 386.06107 112.39056 curveto -385.54544 112.51036 385.06627 112.69005 384.62357 112.92963 curveto -384.62357 110.79681 lineto -385.22252 110.65099 385.82408 110.54161 386.42825 110.46869 curveto -387.03242 110.39057 387.63658 110.35151 388.24075 110.3515 curveto -389.81887 110.35151 390.95689 110.66401 391.65482 111.289 curveto -392.35793 111.9088 392.70949 112.91922 392.7095 114.32025 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -398.38138 108.07806 moveto -398.38138 110.56244 lineto -401.26419 110.56244 lineto -401.26419 112.56244 lineto -398.38138 112.56244 lineto -398.38138 116.27338 lineto -398.38137 116.67963 398.4621 116.95567 398.62357 117.1015 curveto -398.78502 117.24213 399.10533 117.31244 399.5845 117.31244 curveto -401.022 117.31244 lineto -401.022 119.31244 lineto -398.62357 119.31244 lineto -397.5194 119.31244 396.73554 119.08327 396.272 118.62494 curveto -395.81367 118.1614 395.5845 117.37755 395.5845 116.27338 curveto -395.5845 112.56244 lineto -394.19388 112.56244 lineto -394.19388 110.56244 lineto -395.5845 110.56244 lineto -395.5845 108.07806 lineto -398.38138 108.07806 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -411.71732 114.914 moveto -411.71732 115.71088 lineto -405.17825 115.71088 lineto -405.24596 116.36713 405.48294 116.85932 405.88919 117.18744 curveto -406.29544 117.51557 406.86314 117.67963 407.59232 117.67963 curveto -408.18085 117.67963 408.78241 117.59369 409.397 117.42181 curveto -410.01679 117.24473 410.6522 116.97911 411.30325 116.62494 curveto -411.30325 118.78119 lineto -410.64179 119.03119 409.98033 119.21869 409.31888 119.34369 curveto -408.65741 119.4739 407.99596 119.539 407.3345 119.539 curveto -405.75117 119.539 404.5194 119.13796 403.63919 118.33588 curveto -402.76419 117.52859 402.32669 116.39838 402.32669 114.94525 curveto -402.32669 113.51817 402.75638 112.39578 403.61575 111.57806 curveto -404.48033 110.76036 405.66783 110.35151 407.17825 110.3515 curveto -408.55325 110.35151 409.6522 110.76557 410.47513 111.59369 curveto -411.30324 112.42182 411.71731 113.52859 411.71732 114.914 curveto -408.84232 113.98431 moveto -408.84231 113.45307 408.68606 113.02599 408.37357 112.70306 curveto -408.06627 112.37495 407.66262 112.21088 407.16263 112.21088 curveto -406.62096 112.21088 406.18085 112.36453 405.84232 112.67181 curveto -405.50377 112.9739 405.29283 113.4114 405.2095 113.98431 curveto -408.84232 113.98431 lineto -fill -grestore -grestore -0 0 0 setrgbcolor -[] 0 setdash -1 setlinewidth -0 setlinejoin -0 setlinecap -newpath -460.13031 263.40024 moveto -460.13031 289.97505 lineto -stroke -gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 460.13031 289.97505] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -gsave [-0.2378166 0 0 0.2269133 546.17466 132.00569] concat -0 0 0 setrgbcolor -[] 0 setdash -4.3013706 setlinewidth -1 setlinejoin -1 setlinecap -newpath -208.65508 282.05865 moveto -193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto -43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto -45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto -177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto -stroke -gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -grestore -gsave [-0.2378166 0 0 0.2269133 546.17465 87.714359] concat -0 0 0 setrgbcolor -[] 0 setdash -4.3013706 setlinewidth -2 setlinejoin -1 setlinecap -newpath -208.65508 282.05865 moveto -193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto -43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto -45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto -177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto -stroke -gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -grestore -gsave [-0.2378166 0 0 0.2269133 546.17465 176.29703] concat -0 0 0 setrgbcolor -[] 0 setdash -4.3013706 setlinewidth -1 setlinejoin -1 setlinecap -newpath -208.65508 282.05865 moveto -193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto -43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto -45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto -177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto -stroke -gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -grestore -gsave [0 0.2378166 0.2269133 0 399.60191 71.056696] concat -0 0 0 setrgbcolor -[] 0 setdash -4.3013706 setlinewidth -1 setlinejoin -1 setlinecap -newpath -208.65508 282.05865 moveto -193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto -43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto -45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto -177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto -stroke -gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat -gsave -0 0 0 setrgbcolor -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -eofill -grestore -0 0 0 setrgbcolor -[] 0 setdash -1.25 setlinewidth -0 setlinejoin -0 setlinecap -newpath -5.77 0 moveto --2.88 5 lineto --2.88 -5 lineto -5.77 0 lineto -closepath -stroke -grestore -grestore -gsave [1 0 0 1 17.216929 6.5104864] concat -grestore -gsave -0 0 0 setrgbcolor -newpath -187.35507 103.05839 moveto -187.35507 104.61112 lineto -189.20566 104.61112 lineto -189.20566 105.30936 lineto -187.35507 105.30936 lineto -187.35507 108.27811 lineto -187.35507 108.72408 187.41529 109.01054 187.53574 109.13749 curveto -187.65943 109.26444 187.90846 109.32792 188.28281 109.32792 curveto -189.20566 109.32792 lineto -189.20566 110.07987 lineto -188.28281 110.07987 lineto -187.58944 110.07987 187.11093 109.95129 186.84726 109.69413 curveto -186.58359 109.43371 186.45175 108.96171 186.45175 108.27811 curveto -186.45175 105.30936 lineto -185.79257 105.30936 lineto -185.79257 104.61112 lineto -186.45175 104.61112 lineto -186.45175 103.05839 lineto -187.35507 103.05839 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -194.93808 106.77909 moveto -194.93808 110.07987 lineto -194.03964 110.07987 lineto -194.03964 106.80839 lineto -194.03964 106.29081 193.93873 105.90344 193.73691 105.64628 curveto -193.53508 105.38912 193.23235 105.26054 192.8287 105.26054 curveto -192.34368 105.26054 191.96119 105.41516 191.68124 105.7244 curveto -191.40129 106.03365 191.26132 106.4552 191.26132 106.98905 curveto -191.26132 110.07987 lineto -190.358 110.07987 lineto -190.358 102.48222 lineto -191.26132 102.48222 lineto -191.26132 105.46073 lineto -191.47616 105.13196 191.72844 104.88619 192.01816 104.72343 curveto -192.31112 104.56067 192.64804 104.47929 193.0289 104.47929 curveto -193.65715 104.47929 194.13241 104.6746 194.45468 105.06522 curveto -194.77694 105.4526 194.93807 106.02389 194.93808 106.77909 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -201.41757 107.12089 moveto -201.41757 107.56034 lineto -197.28671 107.56034 lineto -197.32577 108.17883 197.51132 108.65084 197.84335 108.97636 curveto -198.17864 109.29862 198.64413 109.45976 199.23984 109.45975 curveto -199.58489 109.45976 199.91854 109.41744 200.24081 109.3328 curveto -200.56633 109.24817 200.8886 109.12121 201.20761 108.95194 curveto -201.20761 109.80155 lineto -200.88534 109.93827 200.55494 110.04244 200.2164 110.11405 curveto -199.87785 110.18567 199.53443 110.22147 199.18613 110.22147 curveto -198.31373 110.22147 197.622 109.96757 197.11093 109.45975 curveto -196.60312 108.95194 196.34921 108.2651 196.34921 107.39921 curveto -196.34921 106.50403 196.5901 105.79439 197.07187 105.2703 curveto -197.55689 104.74296 198.20956 104.47929 199.02988 104.47929 curveto -199.76555 104.47929 200.3466 104.71692 200.77304 105.19218 curveto -201.20272 105.66419 201.41757 106.30709 201.41757 107.12089 curveto -200.51913 106.85722 moveto -200.51262 106.36568 200.37427 105.97343 200.1041 105.68046 curveto -199.83716 105.38749 199.48235 105.24101 199.03964 105.241 curveto -198.53834 105.24101 198.13632 105.38261 197.83359 105.66581 curveto -197.53411 105.94902 197.36158 106.34778 197.31601 106.8621 curveto -200.51913 106.85722 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -205.01132 105.241 moveto -204.52955 105.24101 204.14869 105.42981 203.86874 105.80741 curveto -203.58879 106.18176 203.44882 106.69609 203.44882 107.35038 curveto -203.44882 108.00468 203.58717 108.52063 203.86386 108.89823 curveto -204.14381 109.27258 204.52629 109.45976 205.01132 109.45975 curveto -205.48983 109.45976 205.86907 109.27095 206.14902 108.89335 curveto -206.42896 108.51575 206.56893 108.00142 206.56894 107.35038 curveto -206.56893 106.7026 206.42896 106.1899 206.14902 105.81229 curveto -205.86907 105.43144 205.48983 105.24101 205.01132 105.241 curveto -205.01132 104.47929 moveto -205.79257 104.47929 206.40617 104.7332 206.85214 105.241 curveto -207.2981 105.74882 207.52108 106.45195 207.52109 107.35038 curveto -207.52108 108.24556 207.2981 108.94869 206.85214 109.45975 curveto -206.40617 109.96757 205.79257 110.22147 205.01132 110.22147 curveto -204.22681 110.22147 203.61158 109.96757 203.16562 109.45975 curveto -202.72291 108.94869 202.50156 108.24556 202.50156 107.35038 curveto -202.50156 106.45195 202.72291 105.74882 203.16562 105.241 curveto -203.61158 104.7332 204.22681 104.47929 205.01132 104.47929 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -212.17441 105.45097 moveto -212.07349 105.39238 211.96282 105.35006 211.84238 105.32401 curveto -211.72519 105.29472 211.59498 105.28007 211.45175 105.28007 curveto -210.94394 105.28007 210.55331 105.44609 210.27988 105.77811 curveto -210.00969 106.10689 209.8746 106.58053 209.8746 107.19901 curveto -209.8746 110.07987 lineto -208.97128 110.07987 lineto -208.97128 104.61112 lineto -209.8746 104.61112 lineto -209.8746 105.46073 lineto -210.0634 105.12871 210.30917 104.88294 210.61191 104.72343 curveto -210.91464 104.56067 211.28248 104.47929 211.71542 104.47929 curveto -211.77727 104.47929 211.84563 104.48417 211.9205 104.49393 curveto -211.99537 104.50045 212.07838 104.51184 212.16953 104.52811 curveto -212.17441 105.45097 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -217.58945 107.12089 moveto -217.58945 107.56034 lineto -213.45859 107.56034 lineto -213.49765 108.17883 213.6832 108.65084 214.01523 108.97636 curveto -214.35051 109.29862 214.81601 109.45976 215.41171 109.45975 curveto -215.75676 109.45976 216.09042 109.41744 216.41269 109.3328 curveto -216.73821 109.24817 217.06047 109.12121 217.37949 108.95194 curveto -217.37949 109.80155 lineto -217.05722 109.93827 216.72681 110.04244 216.38828 110.11405 curveto -216.04973 110.18567 215.70631 110.22147 215.358 110.22147 curveto -214.4856 110.22147 213.79387 109.96757 213.28281 109.45975 curveto -212.77499 108.95194 212.52109 108.2651 212.52109 107.39921 curveto -212.52109 106.50403 212.76197 105.79439 213.24374 105.2703 curveto -213.72877 104.74296 214.38144 104.47929 215.20175 104.47929 curveto -215.93742 104.47929 216.51848 104.71692 216.94492 105.19218 curveto -217.3746 105.66419 217.58944 106.30709 217.58945 107.12089 curveto -216.69101 106.85722 moveto -216.68449 106.36568 216.54615 105.97343 216.27597 105.68046 curveto -216.00904 105.38749 215.65422 105.24101 215.21152 105.241 curveto -214.71021 105.24101 214.30819 105.38261 214.00546 105.66581 curveto -213.70598 105.94902 213.53346 106.34778 213.48788 106.8621 curveto -216.69101 106.85722 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -223.32187 105.66093 moveto -223.54647 105.25729 223.81503 104.95943 224.12753 104.76737 curveto -224.44003 104.57532 224.80786 104.47929 225.23105 104.47929 curveto -225.8007 104.47929 226.24016 104.67949 226.54941 105.07987 curveto -226.85864 105.47701 227.01327 106.04342 227.01328 106.77909 curveto -227.01328 110.07987 lineto -226.10995 110.07987 lineto -226.10995 106.80839 lineto -226.10995 106.2843 226.01717 105.89531 225.83163 105.6414 curveto -225.64608 105.38749 225.36288 105.26054 224.98203 105.26054 curveto -224.51652 105.26054 224.14869 105.41516 223.87851 105.7244 curveto -223.60832 106.03365 223.47323 106.4552 223.47324 106.98905 curveto -223.47324 110.07987 lineto -222.56992 110.07987 lineto -222.56992 106.80839 lineto -222.56991 106.28105 222.47714 105.89205 222.2916 105.6414 curveto -222.10604 105.38749 221.81959 105.26054 221.43222 105.26054 curveto -220.97323 105.26054 220.60865 105.41679 220.33847 105.72929 curveto -220.06829 106.03854 219.9332 106.45846 219.9332 106.98905 curveto -219.9332 110.07987 lineto -219.02988 110.07987 lineto -219.02988 104.61112 lineto -219.9332 104.61112 lineto -219.9332 105.46073 lineto -220.13827 105.12545 220.38404 104.87805 220.6705 104.71854 curveto -220.95696 104.55904 221.29713 104.47929 221.69101 104.47929 curveto -222.08814 104.47929 222.42505 104.5802 222.70175 104.78202 curveto -222.98169 104.98385 223.1884 105.27682 223.32187 105.66093 curveto -fill -grestore -gsave [1 0 0 1 17.216929 6.5104864] concat -grestore -gsave -0 0 0 setrgbcolor -newpath -470.46808 277.74594 moveto -470.46808 278.40675 470.60317 278.92596 470.87335 279.30356 curveto -471.14679 279.67791 471.52114 279.86508 471.9964 279.86508 curveto -472.47166 279.86508 472.846 279.67791 473.11945 279.30356 curveto -473.39288 278.92596 473.5296 278.40675 473.5296 277.74594 curveto -473.5296 277.08514 473.39288 276.56756 473.11945 276.19321 curveto -472.846 275.81561 472.47166 275.62681 471.9964 275.6268 curveto -471.52114 275.62681 471.14679 275.81561 470.87335 276.19321 curveto -470.60317 276.56756 470.46808 277.08514 470.46808 277.74594 curveto -473.5296 279.65512 moveto -473.3408 279.98064 473.10154 280.22315 472.81183 280.38266 curveto -472.52537 280.53891 472.18032 280.61703 471.77667 280.61703 curveto -471.11586 280.61703 470.57713 280.35336 470.16046 279.82602 curveto -469.74705 279.29868 469.54034 278.60532 469.54034 277.74594 curveto -469.54034 276.88657 469.74705 276.19321 470.16046 275.66586 curveto -470.57713 275.13852 471.11586 274.87485 471.77667 274.87485 curveto -472.18032 274.87485 472.52537 274.95461 472.81183 275.11411 curveto -473.10154 275.27036 473.3408 275.51125 473.5296 275.83676 curveto -473.5296 275.00668 lineto -474.42804 275.00668 lineto -474.42804 282.55551 lineto -473.5296 282.55551 lineto -473.5296 279.65512 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -480.95636 277.51645 moveto -480.95636 277.9559 lineto -476.8255 277.9559 lineto -476.86456 278.57439 477.05011 279.0464 477.38214 279.37192 curveto -477.71743 279.69418 478.18292 279.85532 478.77863 279.85532 curveto -479.12367 279.85532 479.45733 279.813 479.7796 279.72836 curveto -480.10512 279.64373 480.42738 279.51678 480.7464 279.3475 curveto -480.7464 280.19711 lineto -480.42413 280.33383 480.09372 280.438 479.75519 280.50961 curveto -479.41664 280.58123 479.07322 280.61703 478.72491 280.61703 curveto -477.85252 280.61703 477.16079 280.36313 476.64972 279.85532 curveto -476.14191 279.3475 475.888 278.66066 475.888 277.79477 curveto -475.888 276.89959 476.12889 276.18996 476.61066 275.66586 curveto -477.09568 275.13852 477.74835 274.87485 478.56866 274.87485 curveto -479.30434 274.87485 479.88539 275.11248 480.31183 275.58774 curveto -480.74151 276.05975 480.95635 276.70265 480.95636 277.51645 curveto -480.05792 277.25278 moveto -480.05141 276.76124 479.91306 276.36899 479.64288 276.07602 curveto -479.37595 275.78306 479.02113 275.63657 478.57843 275.63657 curveto -478.07713 275.63657 477.67511 275.77817 477.37238 276.06137 curveto -477.07289 276.34458 476.90037 276.74334 476.8548 277.25766 curveto -480.05792 277.25278 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -486.0296 275.83676 moveto -486.0296 272.87778 lineto -486.92804 272.87778 lineto -486.92804 280.47543 lineto -486.0296 280.47543 lineto -486.0296 279.65512 lineto -485.8408 279.98064 485.60154 280.22315 485.31183 280.38266 curveto -485.02537 280.53891 484.68032 280.61703 484.27667 280.61703 curveto -483.61586 280.61703 483.07713 280.35336 482.66046 279.82602 curveto -482.24705 279.29868 482.04034 278.60532 482.04034 277.74594 curveto -482.04034 276.88657 482.24705 276.19321 482.66046 275.66586 curveto -483.07713 275.13852 483.61586 274.87485 484.27667 274.87485 curveto -484.68032 274.87485 485.02537 274.95461 485.31183 275.11411 curveto -485.60154 275.27036 485.8408 275.51125 486.0296 275.83676 curveto -482.96808 277.74594 moveto -482.96808 278.40675 483.10317 278.92596 483.37335 279.30356 curveto -483.64679 279.67791 484.02114 279.86508 484.4964 279.86508 curveto -484.97166 279.86508 485.346 279.67791 485.61945 279.30356 curveto -485.89288 278.92596 486.0296 278.40675 486.0296 277.74594 curveto -486.0296 277.08514 485.89288 276.56756 485.61945 276.19321 curveto -485.346 275.81561 484.97166 275.62681 484.4964 275.6268 curveto -484.02114 275.62681 483.64679 275.81561 483.37335 276.19321 curveto -483.10317 276.56756 482.96808 277.08514 482.96808 277.74594 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -550.54895 236.85474 moveto -550.54895 237.51555 550.68404 238.03475 550.95422 238.41235 curveto -551.22766 238.7867 551.60201 238.97388 552.07727 238.97388 curveto -552.55253 238.97388 552.92688 238.7867 553.20032 238.41235 curveto -553.47375 238.03475 553.61047 237.51555 553.61047 236.85474 curveto -553.61047 236.19393 553.47375 235.67635 553.20032 235.302 curveto -552.92688 234.9244 552.55253 234.7356 552.07727 234.7356 curveto -551.60201 234.7356 551.22766 234.9244 550.95422 235.302 curveto -550.68404 235.67635 550.54895 236.19393 550.54895 236.85474 curveto -553.61047 238.76392 moveto -553.42167 239.08944 553.18241 239.33195 552.8927 239.49146 curveto -552.60624 239.64771 552.26119 239.72583 551.85754 239.72583 curveto -551.19673 239.72583 550.658 239.46216 550.24133 238.93481 curveto -549.82792 238.40747 549.62122 237.71411 549.62122 236.85474 curveto -549.62122 235.99536 549.82792 235.30201 550.24133 234.77466 curveto -550.658 234.24732 551.19673 233.98365 551.85754 233.98364 curveto -552.26119 233.98365 552.60624 234.0634 552.8927 234.2229 curveto -553.18241 234.37916 553.42167 234.62004 553.61047 234.94556 curveto -553.61047 234.11548 lineto -554.50891 234.11548 lineto -554.50891 241.66431 lineto -553.61047 241.66431 lineto -553.61047 238.76392 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -561.03723 236.62524 moveto -561.03723 237.0647 lineto -556.90637 237.0647 lineto -556.94543 237.68319 557.13098 238.15519 557.46301 238.48071 curveto -557.7983 238.80298 558.26379 238.96411 558.8595 238.96411 curveto -559.20455 238.96411 559.5382 238.92179 559.86047 238.83716 curveto -560.18599 238.75252 560.50825 238.62557 560.82727 238.4563 curveto -560.82727 239.30591 lineto -560.505 239.44263 560.1746 239.54679 559.83606 239.61841 curveto -559.49751 239.69002 559.15409 239.72583 558.80579 239.72583 curveto -557.93339 239.72583 557.24166 239.47192 556.73059 238.96411 curveto -556.22278 238.4563 555.96887 237.76945 555.96887 236.90356 curveto -555.96887 236.00839 556.20976 235.29875 556.69153 234.77466 curveto -557.17655 234.24732 557.82922 233.98365 558.64954 233.98364 curveto -559.38521 233.98365 559.96626 234.22128 560.3927 234.69653 curveto -560.82238 235.16854 561.03723 235.81145 561.03723 236.62524 curveto -560.13879 236.36157 moveto -560.13228 235.87004 559.99393 235.47779 559.72375 235.18481 curveto -559.45682 234.89185 559.10201 234.74537 558.6593 234.74536 curveto -558.158 234.74537 557.75598 234.88697 557.45325 235.17017 curveto -557.15377 235.45337 556.98124 235.85214 556.93567 236.36646 curveto -560.13879 236.36157 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -566.11047 234.94556 moveto -566.11047 231.98657 lineto -567.00891 231.98657 lineto -567.00891 239.58423 lineto -566.11047 239.58423 lineto -566.11047 238.76392 lineto -565.92167 239.08944 565.68241 239.33195 565.3927 239.49146 curveto -565.10624 239.64771 564.76119 239.72583 564.35754 239.72583 curveto -563.69673 239.72583 563.158 239.46216 562.74133 238.93481 curveto -562.32792 238.40747 562.12122 237.71411 562.12122 236.85474 curveto -562.12122 235.99536 562.32792 235.30201 562.74133 234.77466 curveto -563.158 234.24732 563.69673 233.98365 564.35754 233.98364 curveto -564.76119 233.98365 565.10624 234.0634 565.3927 234.2229 curveto -565.68241 234.37916 565.92167 234.62004 566.11047 234.94556 curveto -563.04895 236.85474 moveto -563.04895 237.51555 563.18404 238.03475 563.45422 238.41235 curveto -563.72766 238.7867 564.10201 238.97388 564.57727 238.97388 curveto -565.05253 238.97388 565.42688 238.7867 565.70032 238.41235 curveto -565.97375 238.03475 566.11047 237.51555 566.11047 236.85474 curveto -566.11047 236.19393 565.97375 235.67635 565.70032 235.302 curveto -565.42688 234.9244 565.05253 234.7356 564.57727 234.7356 curveto -564.10201 234.7356 563.72766 234.9244 563.45422 235.302 curveto -563.18404 235.67635 563.04895 236.19393 563.04895 236.85474 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -553.10266 183.66447 moveto -553.10266 184.41154 lineto -552.24329 184.41154 lineto -551.92102 184.41155 551.69641 184.47666 551.56946 184.60686 curveto -551.44576 184.73707 551.38391 184.97145 551.38391 185.30998 curveto -551.38391 185.79338 lineto -552.8634 185.79338 lineto -552.8634 186.49162 lineto -551.38391 186.49162 lineto -551.38391 191.26213 lineto -550.48059 191.26213 lineto -550.48059 186.49162 lineto -549.62122 186.49162 lineto -549.62122 185.79338 lineto -550.48059 185.79338 lineto -550.48059 185.41252 lineto -550.48059 184.8038 550.62219 184.3611 550.9054 184.0844 curveto -551.1886 183.80446 551.63782 183.66448 552.25305 183.66447 curveto -553.10266 183.66447 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -553.84973 185.79338 moveto -554.74817 185.79338 lineto -554.74817 191.26213 lineto -553.84973 191.26213 lineto -553.84973 185.79338 lineto -553.84973 183.66447 moveto -554.74817 183.66447 lineto -554.74817 184.80217 lineto -553.84973 184.80217 lineto -553.84973 183.66447 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -561.16907 185.79338 moveto -559.19153 188.45451 lineto -561.27161 191.26213 lineto -560.21204 191.26213 lineto -558.62024 189.11369 lineto -557.02844 191.26213 lineto -555.96887 191.26213 lineto -558.0929 188.4008 lineto -556.14954 185.79338 lineto -557.20911 185.79338 lineto -558.6593 187.74162 lineto -560.1095 185.79338 lineto -561.16907 185.79338 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -552.81946 198.51311 moveto -552.09354 198.51311 551.59061 198.59612 551.31067 198.76213 curveto -551.03072 198.92815 550.89075 199.21135 550.89075 199.61174 curveto -550.89075 199.93075 550.99491 200.18466 551.20325 200.37346 curveto -551.41483 200.55901 551.70129 200.65178 552.06262 200.65178 curveto -552.56067 200.65178 552.95943 200.476 553.25891 200.12444 curveto -553.56164 199.76962 553.71301 199.29924 553.71301 198.7133 curveto -553.71301 198.51311 lineto -552.81946 198.51311 lineto -554.61145 198.14201 moveto -554.61145 201.26213 lineto -553.71301 201.26213 lineto -553.71301 200.43205 lineto -553.50793 200.76408 553.2524 201.00985 552.94641 201.16936 curveto -552.64042 201.32561 552.26607 201.40373 551.82336 201.40373 curveto -551.26347 201.40373 550.8175 201.24748 550.48547 200.93498 curveto -550.1567 200.61923 549.99231 200.19768 549.99231 199.67033 curveto -549.99231 199.0551 550.19739 198.59123 550.60754 198.27873 curveto -551.02095 197.96624 551.63619 197.80999 552.45325 197.80998 curveto -553.71301 197.80998 lineto -553.71301 197.72209 lineto -553.71301 197.30868 553.57629 196.98967 553.30286 196.76506 curveto -553.03267 196.5372 552.65181 196.42327 552.16028 196.42326 curveto -551.84778 196.42327 551.54341 196.4607 551.24719 196.53557 curveto -550.95097 196.61044 550.66614 196.72275 550.3927 196.87248 curveto -550.3927 196.0424 lineto -550.72147 195.91546 551.04049 195.82106 551.34973 195.7592 curveto -551.65897 195.6941 551.96008 195.66155 552.25305 195.66154 curveto -553.04406 195.66155 553.63488 195.86663 554.02551 196.27678 curveto -554.41613 196.68694 554.61144 197.30868 554.61145 198.14201 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -559.95325 195.95451 moveto -559.95325 196.80412 lineto -559.69934 196.67392 559.43567 196.57626 559.16223 196.51115 curveto -558.88879 196.44605 558.60559 196.4135 558.31262 196.4135 curveto -557.86666 196.4135 557.53137 196.48186 557.30676 196.61858 curveto -557.08541 196.7553 556.97473 196.96038 556.97473 197.23381 curveto -556.97473 197.44215 557.05448 197.60654 557.21399 197.72697 curveto -557.37349 197.84417 557.69413 197.95647 558.1759 198.06389 curveto -558.48352 198.13225 lineto -559.12154 198.26897 559.57401 198.46265 559.84094 198.7133 curveto -560.11112 198.9607 560.24621 199.30738 560.24622 199.75334 curveto -560.24621 200.26116 560.04439 200.66317 559.64075 200.9594 curveto -559.24035 201.25562 558.6886 201.40373 557.98547 201.40373 curveto -557.6925 201.40373 557.38651 201.37444 557.0675 201.31584 curveto -556.75175 201.2605 556.41809 201.17587 556.06653 201.06194 curveto -556.06653 200.1342 lineto -556.39856 200.30673 556.72571 200.43694 557.04797 200.52483 curveto -557.37024 200.60946 557.68925 200.65178 558.005 200.65178 curveto -558.42818 200.65178 558.7537 200.58017 558.98157 200.43694 curveto -559.20943 200.29045 559.32336 200.08537 559.32336 199.8217 curveto -559.32336 199.57756 559.24035 199.39039 559.07434 199.26018 curveto -558.91158 199.12997 558.55188 199.00465 557.99524 198.8842 curveto -557.68274 198.81096 lineto -557.1261 198.69377 556.72408 198.51474 556.47668 198.27385 curveto -556.22929 198.02971 556.10559 197.69605 556.10559 197.27287 curveto -556.10559 196.75855 556.28788 196.36142 556.65247 196.08147 curveto -557.01705 195.80152 557.53463 195.66155 558.2052 195.66154 curveto -558.53723 195.66155 558.84973 195.68596 559.1427 195.73479 curveto -559.43567 195.78362 559.70585 195.85686 559.95325 195.95451 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -565.16809 195.95451 moveto -565.16809 196.80412 lineto -564.91418 196.67392 564.65051 196.57626 564.37708 196.51115 curveto -564.10363 196.44605 563.82043 196.4135 563.52747 196.4135 curveto -563.0815 196.4135 562.74621 196.48186 562.52161 196.61858 curveto -562.30025 196.7553 562.18957 196.96038 562.18958 197.23381 curveto -562.18957 197.44215 562.26933 197.60654 562.42883 197.72697 curveto -562.58834 197.84417 562.90897 197.95647 563.39075 198.06389 curveto -563.69836 198.13225 lineto -564.33638 198.26897 564.78886 198.46265 565.05579 198.7133 curveto -565.32596 198.9607 565.46105 199.30738 565.46106 199.75334 curveto -565.46105 200.26116 565.25923 200.66317 564.85559 200.9594 curveto -564.4552 201.25562 563.90344 201.40373 563.20032 201.40373 curveto -562.90735 201.40373 562.60136 201.37444 562.28235 201.31584 curveto -561.96659 201.2605 561.63293 201.17587 561.28137 201.06194 curveto -561.28137 200.1342 lineto -561.6134 200.30673 561.94055 200.43694 562.26282 200.52483 curveto -562.58508 200.60946 562.90409 200.65178 563.21985 200.65178 curveto -563.64302 200.65178 563.96854 200.58017 564.19641 200.43694 curveto -564.42427 200.29045 564.5382 200.08537 564.53821 199.8217 curveto -564.5382 199.57756 564.4552 199.39039 564.28918 199.26018 curveto -564.12642 199.12997 563.76672 199.00465 563.21008 198.8842 curveto -562.89758 198.81096 lineto -562.34094 198.69377 561.93892 198.51474 561.69153 198.27385 curveto -561.44413 198.02971 561.32043 197.69605 561.32043 197.27287 curveto -561.32043 196.75855 561.50273 196.36142 561.86731 196.08147 curveto -562.23189 195.80152 562.74947 195.66155 563.42004 195.66154 curveto -563.75207 195.66155 564.06457 195.68596 564.35754 195.73479 curveto -564.65051 195.78362 564.92069 195.85686 565.16809 195.95451 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -566.80383 199.10393 moveto -566.80383 195.79338 lineto -567.70227 195.79338 lineto -567.70227 199.06975 lineto -567.70227 199.58733 567.80318 199.97632 568.005 200.23674 curveto -568.20683 200.4939 568.50956 200.62248 568.91321 200.62248 curveto -569.39823 200.62248 569.78072 200.46786 570.06067 200.15862 curveto -570.34387 199.84937 570.48547 199.42782 570.48547 198.89397 curveto -570.48547 195.79338 lineto -571.38391 195.79338 lineto -571.38391 201.26213 lineto -570.48547 201.26213 lineto -570.48547 200.42229 lineto -570.26737 200.75432 570.01346 201.00171 569.72375 201.16447 curveto -569.43729 201.32398 569.10363 201.40373 568.72278 201.40373 curveto -568.09452 201.40373 567.61763 201.20842 567.29211 200.81779 curveto -566.96659 200.42717 566.80383 199.85588 566.80383 199.10393 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -577.50208 196.84319 moveto -577.72668 196.43954 577.99523 196.14169 578.30774 195.94963 curveto -578.62023 195.75758 578.98807 195.66155 579.41125 195.66154 curveto -579.98091 195.66155 580.42036 195.86175 580.72961 196.26213 curveto -581.03885 196.65927 581.19347 197.22568 581.19348 197.96135 curveto -581.19348 201.26213 lineto -580.29016 201.26213 lineto -580.29016 197.99065 lineto -580.29015 197.46656 580.19738 197.07756 580.01184 196.82365 curveto -579.82629 196.56975 579.54308 196.4428 579.16223 196.44279 curveto -578.69673 196.4428 578.32889 196.59742 578.05872 196.90666 curveto -577.78853 197.21591 577.65344 197.63746 577.65344 198.17131 curveto -577.65344 201.26213 lineto -576.75012 201.26213 lineto -576.75012 197.99065 lineto -576.75012 197.46331 576.65734 197.07431 576.4718 196.82365 curveto -576.28625 196.56975 575.99979 196.4428 575.61243 196.44279 curveto -575.15344 196.4428 574.78886 196.59905 574.51868 196.91154 curveto -574.24849 197.22079 574.1134 197.64072 574.1134 198.17131 curveto -574.1134 201.26213 lineto -573.21008 201.26213 lineto -573.21008 195.79338 lineto -574.1134 195.79338 lineto -574.1134 196.64299 lineto -574.31848 196.30771 574.56425 196.06031 574.85071 195.9008 curveto -575.13716 195.7413 575.47733 195.66155 575.87122 195.66154 curveto -576.26835 195.66155 576.60526 195.76246 576.88196 195.96428 curveto -577.1619 196.16611 577.36861 196.45908 577.50208 196.84319 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -587.66809 198.30315 moveto -587.66809 198.7426 lineto -583.53723 198.7426 lineto -583.57629 199.36109 583.76184 199.8331 584.09387 200.15862 curveto -584.42916 200.48088 584.89465 200.64201 585.49036 200.64201 curveto -585.8354 200.64201 586.16906 200.5997 586.49133 200.51506 curveto -586.81685 200.43043 587.13911 200.30347 587.45813 200.1342 curveto -587.45813 200.98381 lineto -587.13586 201.12053 586.80546 201.2247 586.46692 201.29631 curveto -586.12837 201.36792 585.78495 201.40373 585.43665 201.40373 curveto -584.56425 201.40373 583.87252 201.14983 583.36145 200.64201 curveto -582.85364 200.1342 582.59973 199.44735 582.59973 198.58147 curveto -582.59973 197.68629 582.84062 196.97665 583.32239 196.45256 curveto -583.80741 195.92522 584.46008 195.66155 585.2804 195.66154 curveto -586.01607 195.66155 586.59712 195.89918 587.02356 196.37444 curveto -587.45324 196.84645 587.66809 197.48935 587.66809 198.30315 curveto -586.76965 198.03947 moveto -586.76314 197.54794 586.62479 197.15569 586.35461 196.86272 curveto -586.08768 196.56975 585.73287 196.42327 585.29016 196.42326 curveto -584.78886 196.42327 584.38684 196.56487 584.08411 196.84807 curveto -583.78463 197.13128 583.6121 197.53004 583.56653 198.04436 curveto -586.76965 198.03947 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -553.82532 147.89853 moveto -553.82532 148.60165 lineto -553.52258 148.60165 lineto -552.71203 148.60165 552.16841 148.48121 551.89172 148.24033 curveto -551.61828 147.99944 551.48156 147.5193 551.48157 146.7999 curveto -551.48157 145.6329 lineto -551.48156 145.14137 551.39367 144.8012 551.2179 144.6124 curveto -551.04211 144.4236 550.7231 144.3292 550.26086 144.32919 curveto -549.96301 144.32919 lineto -549.96301 143.63095 lineto -550.26086 143.63095 lineto -550.72636 143.63095 551.04537 143.53818 551.2179 143.35263 curveto -551.39367 143.16383 551.48156 142.82692 551.48157 142.34189 curveto -551.48157 141.17001 lineto -551.48156 140.45062 551.61828 139.9721 551.89172 139.73447 curveto -552.16841 139.49359 552.71203 139.37315 553.52258 139.37314 curveto -553.82532 139.37314 lineto -553.82532 140.07138 lineto -553.49329 140.07138 lineto -553.0343 140.07139 552.73482 140.143 552.59485 140.28622 curveto -552.45487 140.42946 552.38488 140.73057 552.38489 141.18954 curveto -552.38489 142.40048 lineto -552.38488 142.91155 552.31001 143.28265 552.16028 143.51376 curveto -552.01379 143.74489 551.76151 143.90114 551.40344 143.98251 curveto -551.76477 144.07041 552.01867 144.22991 552.16516 144.46103 curveto -552.31164 144.69215 552.38488 145.06162 552.38489 145.56943 curveto -552.38489 146.78036 lineto -552.38488 147.23935 552.45487 147.54046 552.59485 147.68369 curveto -552.73482 147.82691 553.0343 147.89853 553.49329 147.89853 curveto -553.82532 147.89853 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -559.51379 147.89853 moveto -559.85559 147.89853 lineto -560.31132 147.89853 560.60754 147.82854 560.74426 147.68857 curveto -560.88423 147.54859 560.95422 147.24586 560.95422 146.78036 curveto -560.95422 145.56943 lineto -560.95422 145.06162 561.02746 144.69215 561.17395 144.46103 curveto -561.32043 144.22991 561.57434 144.07041 561.93567 143.98251 curveto -561.57434 143.90114 561.32043 143.74489 561.17395 143.51376 curveto -561.02746 143.28265 560.95422 142.91155 560.95422 142.40048 curveto -560.95422 141.18954 lineto -560.95422 140.72731 560.88423 140.4262 560.74426 140.28622 curveto -560.60754 140.143 560.31132 140.07139 559.85559 140.07138 curveto -559.51379 140.07138 lineto -559.51379 139.37314 lineto -559.82141 139.37314 lineto -560.63196 139.37315 561.17232 139.49359 561.4425 139.73447 curveto -561.71594 139.9721 561.85266 140.45062 561.85266 141.17001 curveto -561.85266 142.34189 lineto -561.85266 142.82692 561.94055 143.16383 562.11633 143.35263 curveto -562.29211 143.53818 562.61112 143.63095 563.07336 143.63095 curveto -563.3761 143.63095 lineto -563.3761 144.32919 lineto -563.07336 144.32919 lineto -562.61112 144.3292 562.29211 144.4236 562.11633 144.6124 curveto -561.94055 144.8012 561.85266 145.14137 561.85266 145.6329 curveto -561.85266 146.7999 lineto -561.85266 147.5193 561.71594 147.99944 561.4425 148.24033 curveto -561.17232 148.48121 560.63196 148.60165 559.82141 148.60165 curveto -559.51379 148.60165 lineto -559.51379 147.89853 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -554.20129 153.67001 moveto -554.20129 156.97079 lineto -553.30286 156.97079 lineto -553.30286 153.69931 lineto -553.30285 153.18174 553.20194 152.79437 553.00012 152.5372 curveto -552.7983 152.28004 552.49556 152.15146 552.09192 152.15146 curveto -551.60689 152.15146 551.2244 152.30609 550.94446 152.61533 curveto -550.66451 152.92457 550.52453 153.34612 550.52454 153.87997 curveto -550.52454 156.97079 lineto -549.62122 156.97079 lineto -549.62122 151.50204 lineto -550.52454 151.50204 lineto -550.52454 152.35165 lineto -550.73938 152.02288 550.99166 151.77711 551.28137 151.61435 curveto -551.57434 151.45159 551.91125 151.37021 552.29211 151.37021 curveto -552.92037 151.37021 553.39563 151.56553 553.7179 151.95615 curveto -554.04016 152.34352 554.20129 152.91481 554.20129 153.67001 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -560.68079 154.01181 moveto -560.68079 154.45126 lineto -556.54993 154.45126 lineto -556.58899 155.06975 556.77453 155.54176 557.10657 155.86728 curveto -557.44185 156.18955 557.90735 156.35068 558.50305 156.35068 curveto -558.8481 156.35068 559.18176 156.30836 559.50403 156.22372 curveto -559.82954 156.13909 560.15181 156.01214 560.47083 155.84286 curveto -560.47083 156.69247 lineto -560.14855 156.82919 559.81815 156.93336 559.47961 157.00497 curveto -559.14107 157.07659 558.79764 157.1124 558.44934 157.1124 curveto -557.57694 157.1124 556.88521 156.85849 556.37415 156.35068 curveto -555.86633 155.84287 555.61243 155.15602 555.61243 154.29013 curveto -555.61243 153.39495 555.85331 152.68532 556.33508 152.16122 curveto -556.82011 151.63389 557.47278 151.37021 558.29309 151.37021 curveto -559.02876 151.37021 559.60982 151.60784 560.03625 152.0831 curveto -560.46594 152.55511 560.68078 153.19801 560.68079 154.01181 curveto -559.78235 153.74814 moveto -559.77583 153.25661 559.63749 152.86435 559.36731 152.57138 curveto -559.10038 152.27842 558.74556 152.13193 558.30286 152.13193 curveto -557.80155 152.13193 557.39953 152.27353 557.0968 152.55673 curveto -556.79732 152.83994 556.62479 153.2387 556.57922 153.75302 curveto -559.78235 153.74814 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -566.52551 151.50204 moveto -564.54797 154.16318 lineto -566.62805 156.97079 lineto -565.56848 156.97079 lineto -563.97668 154.82236 lineto -562.38489 156.97079 lineto -561.32532 156.97079 lineto -563.44934 154.10947 lineto -561.50598 151.50204 lineto -562.56555 151.50204 lineto -564.01575 153.45029 lineto -565.46594 151.50204 lineto -566.52551 151.50204 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -568.78625 149.94931 moveto -568.78625 151.50204 lineto -570.63684 151.50204 lineto -570.63684 152.20029 lineto -568.78625 152.20029 lineto -568.78625 155.16904 lineto -568.78625 155.615 568.84647 155.90146 568.96692 156.02841 curveto -569.09061 156.15537 569.33964 156.21884 569.71399 156.21884 curveto -570.63684 156.21884 lineto -570.63684 156.97079 lineto -569.71399 156.97079 lineto -569.02063 156.97079 568.54211 156.84221 568.27844 156.58505 curveto -568.01477 156.32464 567.88293 155.85263 567.88293 155.16904 curveto -567.88293 152.20029 lineto -567.22375 152.20029 lineto -567.22375 151.50204 lineto -567.88293 151.50204 lineto -567.88293 149.94931 lineto -568.78625 149.94931 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -483.33514 94.963516 moveto -483.33514 98.264297 lineto -482.43671 98.264297 lineto -482.43671 94.992813 lineto -482.4367 94.475239 482.33579 94.087869 482.13397 93.830704 curveto -481.93215 93.573547 481.62941 93.444966 481.22577 93.444962 curveto -480.74074 93.444966 480.35825 93.599589 480.07831 93.908829 curveto -479.79836 94.218078 479.65838 94.639627 479.65839 95.173477 curveto -479.65839 98.264297 lineto -478.75507 98.264297 lineto -478.75507 92.795547 lineto -479.65839 92.795547 lineto -479.65839 93.645157 lineto -479.87323 93.316386 480.12551 93.070618 480.41522 92.907852 curveto -480.70819 92.745097 481.0451 92.663717 481.42596 92.663712 curveto -482.05422 92.663717 482.52948 92.859029 482.85175 93.249649 curveto -483.17401 93.637023 483.33514 94.208312 483.33514 94.963516 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -487.25604 93.42543 moveto -486.77427 93.425435 486.39341 93.614237 486.11346 93.991837 curveto -485.83351 94.366189 485.69354 94.880512 485.69354 95.534805 curveto -485.69354 96.189104 485.83189 96.705054 486.10858 97.082657 curveto -486.38853 97.457007 486.77101 97.644181 487.25604 97.64418 curveto -487.73455 97.644181 488.11379 97.455379 488.39374 97.077774 curveto -488.67368 96.700171 488.81366 96.185849 488.81366 95.534805 curveto -488.81366 94.887022 488.67368 94.374327 488.39374 93.996719 curveto -488.11379 93.615865 487.73455 93.425435 487.25604 93.42543 curveto -487.25604 92.663712 moveto -488.03729 92.663717 488.65089 92.917623 489.09686 93.42543 curveto -489.54282 93.933247 489.7658 94.636371 489.76581 95.534805 curveto -489.7658 96.429989 489.54282 97.133114 489.09686 97.64418 curveto -488.65089 98.151993 488.03729 98.405899 487.25604 98.405899 curveto -486.47153 98.405899 485.8563 98.151993 485.41034 97.64418 curveto -484.96763 97.133114 484.74628 96.429989 484.74628 95.534805 curveto -484.74628 94.636371 484.96763 93.933247 485.41034 93.42543 curveto -485.8563 92.917623 486.47153 92.663717 487.25604 92.663712 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -492.13885 91.242813 moveto -492.13885 92.795547 lineto -493.98944 92.795547 lineto -493.98944 93.49379 lineto -492.13885 93.49379 lineto -492.13885 96.46254 lineto -492.13885 96.908505 492.19907 97.194963 492.31952 97.321915 curveto -492.44321 97.448869 492.69224 97.512345 493.06659 97.512344 curveto -493.98944 97.512344 lineto -493.98944 98.264297 lineto -493.06659 98.264297 lineto -492.37323 98.264297 491.89471 98.135717 491.63104 97.878555 curveto -491.36737 97.618139 491.23553 97.146135 491.23553 96.46254 curveto -491.23553 93.49379 lineto -490.57635 93.49379 lineto -490.57635 92.795547 lineto -491.23553 92.795547 lineto -491.23553 91.242813 lineto -492.13885 91.242813 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -499.8537 95.305313 moveto -499.8537 95.744766 lineto -495.72284 95.744766 lineto -495.7619 96.363258 495.94745 96.835262 496.27948 97.160782 curveto -496.61476 97.483048 497.08026 97.644181 497.67596 97.64418 curveto -498.02101 97.644181 498.35467 97.601863 498.67694 97.517227 curveto -499.00246 97.432593 499.32472 97.30564 499.64374 97.136368 curveto -499.64374 97.985977 lineto -499.32147 98.122696 498.99106 98.226863 498.65253 98.298477 curveto -498.31398 98.370092 497.97056 98.405899 497.62225 98.405899 curveto -496.74986 98.405899 496.05812 98.151993 495.54706 97.64418 curveto -495.03924 97.136369 494.78534 96.449521 494.78534 95.583633 curveto -494.78534 94.688455 495.02622 93.97882 495.508 93.454727 curveto -495.99302 92.927389 496.64569 92.663717 497.466 92.663712 curveto -498.20168 92.663717 498.78273 92.901347 499.20917 93.376602 curveto -499.63885 93.848612 499.85369 94.491515 499.8537 95.305313 curveto -498.95526 95.041641 moveto -498.94875 94.550108 498.8104 94.157856 498.54022 93.864883 curveto -498.27329 93.571919 497.91847 93.425435 497.47577 93.42543 curveto -496.97446 93.425435 496.57245 93.567037 496.26971 93.850235 curveto -495.97023 94.133442 495.79771 94.532205 495.75214 95.046524 curveto -498.95526 95.041641 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -478.78925 100.66664 moveto -479.68768 100.66664 lineto -479.68768 108.2643 lineto -478.78925 108.2643 lineto -478.78925 100.66664 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -486.24042 105.30531 moveto -486.24042 105.74477 lineto -482.10956 105.74477 lineto -482.14862 106.36326 482.33417 106.83526 482.6662 107.16078 curveto -483.00148 107.48305 483.46698 107.64418 484.06268 107.64418 curveto -484.40773 107.64418 484.74139 107.60186 485.06366 107.51723 curveto -485.38918 107.43259 485.71144 107.30564 486.03046 107.13637 curveto -486.03046 107.98598 lineto -485.70819 108.1227 485.37778 108.22686 485.03925 108.29848 curveto -484.7007 108.37009 484.35728 108.4059 484.00897 108.4059 curveto -483.13657 108.4059 482.44484 108.15199 481.93378 107.64418 curveto -481.42596 107.13637 481.17206 106.44952 481.17206 105.58363 curveto -481.17206 104.68845 481.41294 103.97882 481.89471 103.45473 curveto -482.37974 102.92739 483.03241 102.66372 483.85272 102.66371 curveto -484.5884 102.66372 485.16945 102.90135 485.59589 103.3766 curveto -486.02557 103.84861 486.24041 104.49151 486.24042 105.30531 curveto -485.34198 105.04164 moveto -485.33546 104.55011 485.19712 104.15786 484.92694 103.86488 curveto -484.66001 103.57192 484.30519 103.42544 483.86249 103.42543 curveto -483.36118 103.42544 482.95917 103.56704 482.65643 103.85023 curveto -482.35695 104.13344 482.18443 104.5322 482.13885 105.04652 curveto -485.34198 105.04164 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -488.6037 101.24281 moveto -488.6037 102.79555 lineto -490.45428 102.79555 lineto -490.45428 103.49379 lineto -488.6037 103.49379 lineto -488.6037 106.46254 lineto -488.6037 106.9085 488.66392 107.19496 488.78436 107.32191 curveto -488.90806 107.44887 489.15708 107.51235 489.53143 107.51234 curveto -490.45428 107.51234 lineto -490.45428 108.2643 lineto -489.53143 108.2643 lineto -488.83807 108.2643 488.35956 108.13572 488.09589 107.87856 curveto -487.83221 107.61814 487.70038 107.14613 487.70038 106.46254 curveto -487.70038 103.49379 lineto -487.0412 103.49379 lineto -487.0412 102.79555 lineto -487.70038 102.79555 lineto -487.70038 101.24281 lineto -488.6037 101.24281 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -44.641342 188.13469 moveto -44.641342 184.82414 lineto -45.53978 184.82414 lineto -45.53978 188.10051 lineto -45.539778 188.61809 45.640689 189.00709 45.842514 189.2675 curveto -46.044335 189.52466 46.347069 189.65324 46.750717 189.65324 curveto -47.23574 189.65324 47.618226 189.49862 47.898178 189.18938 curveto -48.181377 188.88013 48.322978 188.45858 48.322983 187.92473 curveto -48.322983 184.82414 lineto -49.22142 184.82414 lineto -49.22142 190.29289 lineto -48.322983 190.29289 lineto -48.322983 189.45305 lineto -48.10488 189.78508 47.850974 190.03248 47.561264 190.19524 curveto -47.274802 190.35474 46.941144 190.43449 46.560287 190.43449 curveto -45.93203 190.43449 45.455143 190.23918 45.129623 189.84856 curveto -44.804102 189.45793 44.641341 188.88664 44.641342 188.13469 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -54.5681 184.98528 moveto -54.5681 185.83488 lineto -54.31419 185.70468 54.050518 185.60702 53.777084 185.54192 curveto -53.503643 185.47682 53.220441 185.44426 52.927475 185.44426 curveto -52.481509 185.44426 52.146223 185.51262 51.921616 185.64934 curveto -51.70026 185.78606 51.589583 185.99114 51.589584 186.26457 curveto -51.589583 186.47291 51.669335 186.6373 51.828842 186.75774 curveto -51.988346 186.87493 52.308983 186.98723 52.790756 187.09465 curveto -53.098373 187.16301 lineto -53.736391 187.29973 54.188864 187.49342 54.455795 187.74406 curveto -54.725973 187.99146 54.861064 188.33814 54.861069 188.7841 curveto -54.861064 189.29192 54.659241 189.69393 54.2556 189.99016 curveto -53.855206 190.28638 53.303448 190.43449 52.600327 190.43449 curveto -52.307356 190.43449 52.001366 190.4052 51.682358 190.3466 curveto -51.366601 190.29126 51.032943 190.20663 50.681381 190.0927 curveto -50.681381 189.16496 lineto -51.013412 189.33749 51.34056 189.4677 51.662827 189.55559 curveto -51.98509 189.64022 52.3041 189.68254 52.619858 189.68254 curveto -53.043032 189.68254 53.368552 189.61093 53.59642 189.4677 curveto -53.824281 189.32121 53.938213 189.11614 53.938217 188.85246 curveto -53.938213 188.60832 53.855206 188.42115 53.689194 188.29094 curveto -53.52643 188.16073 53.16673 188.03541 52.610092 187.91496 curveto -52.297592 187.84172 lineto -51.74095 187.72454 51.338932 187.5455 51.091537 187.30461 curveto -50.844141 187.06047 50.720443 186.72682 50.720444 186.30363 curveto -50.720443 185.78932 50.902735 185.39218 51.267319 185.11223 curveto -51.631901 184.83229 52.149478 184.69231 52.820053 184.69231 curveto -53.152081 184.69231 53.464581 184.71673 53.757553 184.76555 curveto -54.050518 184.81438 54.3207 184.88762 54.5681 184.98528 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -56.296616 184.82414 moveto -57.195053 184.82414 lineto -57.195053 190.29289 lineto -56.296616 190.29289 lineto -56.296616 184.82414 lineto -56.296616 182.69524 moveto -57.195053 182.69524 lineto -57.195053 183.83293 lineto -56.296616 183.83293 lineto -56.296616 182.69524 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -63.615952 186.99211 moveto -63.615952 190.29289 lineto -62.717514 190.29289 lineto -62.717514 187.02141 lineto -62.717509 186.50383 62.616598 186.11646 62.41478 185.8593 curveto -62.212953 185.60214 61.910219 185.47356 61.506577 185.47356 curveto -61.021548 185.47356 60.639061 185.62818 60.359116 185.93742 curveto -60.079166 186.24667 59.939192 186.66822 59.939194 187.20207 curveto -59.939194 190.29289 lineto -59.035873 190.29289 lineto -59.035873 184.82414 lineto -59.939194 184.82414 lineto -59.939194 185.67375 lineto -60.154035 185.34498 60.406314 185.09921 60.69603 184.93645 curveto -60.988996 184.77369 61.325909 184.69231 61.706772 184.69231 curveto -62.335023 184.69231 62.810283 184.88762 63.132553 185.27824 curveto -63.454813 185.66562 63.615946 186.23691 63.615952 186.99211 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -69.016342 187.49504 moveto -69.016338 186.844 68.881247 186.33945 68.611069 185.98137 curveto -68.344138 185.6233 67.968162 185.44426 67.483139 185.44426 curveto -67.001366 185.44426 66.625389 185.6233 66.355209 185.98137 curveto -66.088281 186.33945 65.954817 186.844 65.954819 187.49504 curveto -65.954817 188.14283 66.088281 188.64576 66.355209 189.00383 curveto -66.625389 189.3619 67.001366 189.54094 67.483139 189.54094 curveto -67.968162 189.54094 68.344138 189.3619 68.611069 189.00383 curveto -68.881247 188.64576 69.016338 188.14283 69.016342 187.49504 curveto -69.91478 189.61418 moveto -69.914774 190.54517 69.708069 191.2369 69.294662 191.68938 curveto -68.881247 192.1451 68.248109 192.37297 67.395248 192.37297 curveto -67.079491 192.37297 66.781639 192.34855 66.501694 192.29973 curveto -66.221744 192.25415 65.949934 192.18254 65.686264 192.08488 curveto -65.686264 191.21086 lineto -65.949934 191.35409 66.210351 191.45988 66.467514 191.52824 curveto -66.724673 191.5966 66.986717 191.63078 67.253647 191.63078 curveto -67.842836 191.63078 68.283916 191.47616 68.576889 191.16692 curveto -68.869853 190.86093 69.016338 190.39706 69.016342 189.77531 curveto -69.016342 189.33098 lineto -68.830791 189.65324 68.593161 189.89413 68.303452 190.05363 curveto -68.013734 190.21314 67.667055 190.29289 67.263412 190.29289 curveto -66.592837 190.29289 66.052473 190.03736 65.642319 189.52629 curveto -65.232162 189.01522 65.027084 188.33814 65.027084 187.49504 curveto -65.027084 186.64869 65.232162 185.96998 65.642319 185.45891 curveto -66.052473 184.94785 66.592837 184.69231 67.263412 184.69231 curveto -67.667055 184.69231 68.013734 184.77206 68.303452 184.93156 curveto -68.593161 185.09107 68.830791 185.33196 69.016342 185.65422 curveto -69.016342 184.82414 lineto -69.91478 184.82414 lineto -69.91478 189.61418 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -44.641342 198.13469 moveto -44.641342 194.82414 lineto -45.53978 194.82414 lineto -45.53978 198.10051 lineto -45.539778 198.61809 45.640689 199.00709 45.842514 199.2675 curveto -46.044335 199.52466 46.347069 199.65324 46.750717 199.65324 curveto -47.23574 199.65324 47.618226 199.49862 47.898178 199.18938 curveto -48.181377 198.88013 48.322978 198.45858 48.322983 197.92473 curveto -48.322983 194.82414 lineto -49.22142 194.82414 lineto -49.22142 200.29289 lineto -48.322983 200.29289 lineto -48.322983 199.45305 lineto -48.10488 199.78508 47.850974 200.03248 47.561264 200.19524 curveto -47.274802 200.35474 46.941144 200.43449 46.560287 200.43449 curveto -45.93203 200.43449 45.455143 200.23918 45.129623 199.84856 curveto -44.804102 199.45793 44.641341 198.88664 44.641342 198.13469 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -55.62767 196.99211 moveto -55.62767 200.29289 lineto -54.729233 200.29289 lineto -54.729233 197.02141 lineto -54.729228 196.50383 54.628317 196.11646 54.426498 195.8593 curveto -54.224671 195.60214 53.921937 195.47356 53.518295 195.47356 curveto -53.033266 195.47356 52.65078 195.62818 52.370834 195.93742 curveto -52.090884 196.24667 51.950911 196.66822 51.950912 197.20207 curveto -51.950912 200.29289 lineto -51.047592 200.29289 lineto -51.047592 194.82414 lineto -51.950912 194.82414 lineto -51.950912 195.67375 lineto -52.165754 195.34498 52.418033 195.09921 52.707748 194.93645 curveto -53.000714 194.77369 53.337628 194.69231 53.718491 194.69231 curveto -54.346742 194.69231 54.822002 194.88762 55.144272 195.27824 curveto -55.466532 195.66562 55.627665 196.23691 55.62767 196.99211 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -60.197983 192.69524 moveto -60.197983 193.44231 lineto -59.338608 193.44231 lineto -59.01634 193.44231 58.79173 193.50742 58.66478 193.63762 curveto -58.54108 193.76783 58.479231 194.00221 58.479233 194.34074 curveto -58.479233 194.82414 lineto -59.958725 194.82414 lineto -59.958725 195.52238 lineto -58.479233 195.52238 lineto -58.479233 200.29289 lineto -57.575912 200.29289 lineto -57.575912 195.52238 lineto -56.716537 195.52238 lineto -56.716537 194.82414 lineto -57.575912 194.82414 lineto -57.575912 194.44328 lineto -57.575911 193.83457 57.717513 193.39186 58.000717 193.11516 curveto -58.283918 192.83522 58.733137 192.69524 59.348373 192.69524 curveto -60.197983 192.69524 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -63.064194 195.45403 moveto -62.58242 195.45403 62.201561 195.64283 61.921616 196.02043 curveto -61.641666 196.39478 61.501692 196.90911 61.501694 197.5634 curveto -61.501692 198.2177 61.640038 198.73365 61.916733 199.11125 curveto -62.196679 199.4856 62.579165 199.67278 63.064194 199.67278 curveto -63.542706 199.67278 63.921937 199.48397 64.201889 199.10637 curveto -64.481832 198.72877 64.621806 198.21444 64.621811 197.5634 curveto -64.621806 196.91562 64.481832 196.40292 64.201889 196.02531 curveto -63.921937 195.64446 63.542706 195.45403 63.064194 195.45403 curveto -63.064194 194.69231 moveto -63.84544 194.69231 64.459046 194.94622 64.905014 195.45403 curveto -65.350972 195.96184 65.573954 196.66497 65.573959 197.5634 curveto -65.573954 198.45858 65.350972 199.16171 64.905014 199.67278 curveto -64.459046 200.18059 63.84544 200.43449 63.064194 200.43449 curveto -62.279686 200.43449 61.664452 200.18059 61.218491 199.67278 curveto -60.775781 199.16171 60.554428 198.45858 60.554428 197.5634 curveto -60.554428 196.66497 60.775781 195.96184 61.218491 195.45403 curveto -61.664452 194.94622 62.279686 194.69231 63.064194 194.69231 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -67.058334 192.69524 moveto -67.956772 192.69524 lineto -67.956772 200.29289 lineto -67.058334 200.29289 lineto -67.058334 192.69524 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -73.430405 195.65422 moveto -73.430405 192.69524 lineto -74.328842 192.69524 lineto -74.328842 200.29289 lineto -73.430405 200.29289 lineto -73.430405 199.47258 lineto -73.241598 199.7981 73.002341 200.04061 72.712631 200.20012 curveto -72.426169 200.35637 72.081118 200.43449 71.677475 200.43449 curveto -71.016666 200.43449 70.477929 200.17082 70.061264 199.64348 curveto -69.647852 199.11614 69.441146 198.42278 69.441147 197.5634 curveto -69.441146 196.70403 69.647852 196.01067 70.061264 195.48332 curveto -70.477929 194.95598 71.016666 194.69231 71.677475 194.69231 curveto -72.081118 194.69231 72.426169 194.77206 72.712631 194.93156 curveto -73.002341 195.08782 73.241598 195.3287 73.430405 195.65422 curveto -70.368881 197.5634 moveto -70.36888 198.22421 70.503971 198.74341 70.774155 199.12102 curveto -71.04759 199.49537 71.421939 199.68254 71.897202 199.68254 curveto -72.372458 199.68254 72.746807 199.49537 73.020248 199.12102 curveto -73.293682 198.74341 73.4304 198.22421 73.430405 197.5634 curveto -73.4304 196.9026 73.293682 196.38502 73.020248 196.01067 curveto -72.746807 195.63307 72.372458 195.44426 71.897202 195.44426 curveto -71.421939 195.44426 71.04759 195.63307 70.774155 196.01067 curveto -70.503971 196.38502 70.36888 196.9026 70.368881 197.5634 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -76.179428 194.82414 moveto -77.077866 194.82414 lineto -77.077866 200.29289 lineto -76.179428 200.29289 lineto -76.179428 194.82414 lineto -76.179428 192.69524 moveto -77.077866 192.69524 lineto -77.077866 193.83293 lineto -76.179428 193.83293 lineto -76.179428 192.69524 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -83.498764 196.99211 moveto -83.498764 200.29289 lineto -82.600327 200.29289 lineto -82.600327 197.02141 lineto -82.600322 196.50383 82.499411 196.11646 82.297592 195.8593 curveto -82.095765 195.60214 81.793031 195.47356 81.389389 195.47356 curveto -80.90436 195.47356 80.521874 195.62818 80.241928 195.93742 curveto -79.961978 196.24667 79.822004 196.66822 79.822006 197.20207 curveto -79.822006 200.29289 lineto -78.918686 200.29289 lineto -78.918686 194.82414 lineto -79.822006 194.82414 lineto -79.822006 195.67375 lineto -80.036848 195.34498 80.289126 195.09921 80.578842 194.93645 curveto -80.871808 194.77369 81.208722 194.69231 81.589584 194.69231 curveto -82.217835 194.69231 82.693095 194.88762 83.015366 195.27824 curveto -83.337626 195.66562 83.498759 196.23691 83.498764 196.99211 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -88.899155 197.49504 moveto -88.89915 196.844 88.764059 196.33945 88.493881 195.98137 curveto -88.22695 195.6233 87.850974 195.44426 87.365952 195.44426 curveto -86.884178 195.44426 86.508202 195.6233 86.238022 195.98137 curveto -85.971093 196.33945 85.83763 196.844 85.837631 197.49504 curveto -85.83763 198.14283 85.971093 198.64576 86.238022 199.00383 curveto -86.508202 199.3619 86.884178 199.54094 87.365952 199.54094 curveto -87.850974 199.54094 88.22695 199.3619 88.493881 199.00383 curveto -88.764059 198.64576 88.89915 198.14283 88.899155 197.49504 curveto -89.797592 199.61418 moveto -89.797587 200.54517 89.590881 201.2369 89.177475 201.68938 curveto -88.764059 202.1451 88.130922 202.37297 87.278061 202.37297 curveto -86.962303 202.37297 86.664452 202.34855 86.384506 202.29973 curveto -86.104557 202.25415 85.832747 202.18254 85.569077 202.08488 curveto -85.569077 201.21086 lineto -85.832747 201.35409 86.093163 201.45988 86.350327 201.52824 curveto -86.607486 201.5966 86.86953 201.63078 87.136459 201.63078 curveto -87.725649 201.63078 88.166729 201.47616 88.459702 201.16692 curveto -88.752666 200.86093 88.89915 200.39706 88.899155 199.77531 curveto -88.899155 199.33098 lineto -88.713603 199.65324 88.475973 199.89413 88.186264 200.05363 curveto -87.896547 200.21314 87.549868 200.29289 87.146225 200.29289 curveto -86.47565 200.29289 85.935286 200.03736 85.525131 199.52629 curveto -85.114974 199.01522 84.909896 198.33814 84.909897 197.49504 curveto -84.909896 196.64869 85.114974 195.96998 85.525131 195.45891 curveto -85.935286 194.94785 86.47565 194.69231 87.146225 194.69231 curveto -87.549868 194.69231 87.896547 194.77206 88.186264 194.93156 curveto -88.475973 195.09107 88.713603 195.33196 88.899155 195.65422 curveto -88.899155 194.82414 lineto -89.797592 194.82414 lineto -89.797592 199.61418 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -380.48996 223.50369 moveto -380.48996 225.05643 lineto -382.34055 225.05643 lineto -382.34055 225.75467 lineto -380.48996 225.75467 lineto -380.48996 228.72342 lineto -380.48996 229.16938 380.55018 229.45584 380.67062 229.58279 curveto -380.79432 229.70975 381.04334 229.77322 381.41769 229.77322 curveto -382.34055 229.77322 lineto -382.34055 230.52518 lineto -381.41769 230.52518 lineto -380.72433 230.52518 380.24582 230.3966 379.98215 230.13943 curveto -379.71847 229.87902 379.58664 229.40701 379.58664 228.72342 curveto -379.58664 225.75467 lineto -378.92746 225.75467 lineto -378.92746 225.05643 lineto -379.58664 225.05643 lineto -379.58664 223.50369 lineto -380.48996 223.50369 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -388.07297 227.2244 moveto -388.07297 230.52518 lineto -387.17453 230.52518 lineto -387.17453 227.25369 lineto -387.17453 226.73612 387.07361 226.34875 386.8718 226.09158 curveto -386.66997 225.83443 386.36723 225.70585 385.96359 225.70584 curveto -385.47856 225.70585 385.09608 225.86047 384.81613 226.16971 curveto -384.53618 226.47896 384.39621 226.90051 384.39621 227.43436 curveto -384.39621 230.52518 lineto -383.49289 230.52518 lineto -383.49289 222.92752 lineto -384.39621 222.92752 lineto -384.39621 225.90604 lineto -384.61105 225.57727 384.86333 225.3315 385.15305 225.16873 curveto -385.44601 225.00598 385.78293 224.9246 386.16379 224.92459 curveto -386.79204 224.9246 387.2673 225.11991 387.58957 225.51053 curveto -387.91183 225.8979 388.07296 226.46919 388.07297 227.2244 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -394.55246 227.56619 moveto -394.55246 228.00565 lineto -390.4216 228.00565 lineto -390.46066 228.62414 390.64621 229.09614 390.97824 229.42166 curveto -391.31353 229.74393 391.77902 229.90506 392.37473 229.90506 curveto -392.71977 229.90506 393.05343 229.86274 393.3757 229.77811 curveto -393.70122 229.69347 394.02348 229.56652 394.3425 229.39725 curveto -394.3425 230.24686 lineto -394.02023 230.38358 393.68982 230.48774 393.35129 230.55936 curveto -393.01274 230.63097 392.66932 230.66678 392.32101 230.66678 curveto -391.44862 230.66678 390.75688 230.41287 390.24582 229.90506 curveto -389.73801 229.39725 389.4841 228.7104 389.4841 227.84451 curveto -389.4841 226.94933 389.72498 226.2397 390.20676 225.71561 curveto -390.69178 225.18827 391.34445 224.9246 392.16476 224.92459 curveto -392.90044 224.9246 393.48149 225.16223 393.90793 225.63748 curveto -394.33761 226.10949 394.55245 226.75239 394.55246 227.56619 curveto -393.65402 227.30252 moveto -393.64751 226.81099 393.50916 226.41874 393.23898 226.12576 curveto -392.97205 225.8328 392.61723 225.68631 392.17453 225.68631 curveto -391.67323 225.68631 391.27121 225.82792 390.96848 226.11111 curveto -390.66899 226.39432 390.49647 226.79308 390.4509 227.3074 curveto -393.65402 227.30252 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -400.57297 227.2244 moveto -400.57297 230.52518 lineto -399.67453 230.52518 lineto -399.67453 227.25369 lineto -399.67453 226.73612 399.57361 226.34875 399.3718 226.09158 curveto -399.16997 225.83443 398.86723 225.70585 398.46359 225.70584 curveto -397.97856 225.70585 397.59608 225.86047 397.31613 226.16971 curveto -397.03618 226.47896 396.89621 226.90051 396.89621 227.43436 curveto -396.89621 230.52518 lineto -395.99289 230.52518 lineto -395.99289 225.05643 lineto -396.89621 225.05643 lineto -396.89621 225.90604 lineto -397.11105 225.57727 397.36333 225.3315 397.65305 225.16873 curveto -397.94601 225.00598 398.28293 224.9246 398.66379 224.92459 curveto -399.29204 224.9246 399.7673 225.11991 400.08957 225.51053 curveto -400.41183 225.8979 400.57296 226.46919 400.57297 227.2244 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -239.47623 229.75269 moveto -239.47623 233.05347 lineto -238.57779 233.05347 lineto -238.57779 229.78198 lineto -238.57778 229.26441 238.47687 228.87704 238.27505 228.61987 curveto -238.07323 228.36272 237.77049 228.23414 237.36685 228.23413 curveto -236.88182 228.23414 236.49934 228.38876 236.21939 228.698 curveto -235.93944 229.00725 235.79947 229.4288 235.79947 229.96265 curveto -235.79947 233.05347 lineto -234.89615 233.05347 lineto -234.89615 225.45581 lineto -235.79947 225.45581 lineto -235.79947 228.43433 lineto -236.01431 228.10556 236.26659 227.85979 236.5563 227.69702 curveto -236.84927 227.53427 237.18618 227.45289 237.56705 227.45288 curveto -238.1953 227.45289 238.67056 227.6482 238.99283 228.03882 curveto -239.31509 228.42619 239.47622 228.99748 239.47623 229.75269 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -243.76334 230.30444 moveto -243.03742 230.30445 242.53449 230.38745 242.25455 230.55347 curveto -241.9746 230.71948 241.83462 231.00269 241.83463 231.40308 curveto -241.83462 231.72209 241.93879 231.97599 242.14713 232.16479 curveto -242.35871 232.35034 242.64517 232.44312 243.0065 232.44312 curveto -243.50454 232.44312 243.90331 232.26733 244.20279 231.91577 curveto -244.50552 231.56096 244.65689 231.09058 244.65689 230.50464 curveto -244.65689 230.30444 lineto -243.76334 230.30444 lineto -245.55533 229.93335 moveto -245.55533 233.05347 lineto -244.65689 233.05347 lineto -244.65689 232.22339 lineto -244.45181 232.55542 244.19628 232.80119 243.89029 232.96069 curveto -243.5843 233.11694 243.20995 233.19507 242.76724 233.19507 curveto -242.20734 233.19507 241.76138 233.03882 241.42935 232.72632 curveto -241.10057 232.41056 240.93619 231.98901 240.93619 231.46167 curveto -240.93619 230.84644 241.14127 230.38257 241.55142 230.07007 curveto -241.96483 229.75757 242.58007 229.60132 243.39713 229.60132 curveto -244.65689 229.60132 lineto -244.65689 229.51343 lineto -244.65689 229.10002 244.52017 228.78101 244.24673 228.5564 curveto -243.97655 228.32854 243.59569 228.2146 243.10416 228.2146 curveto -242.79165 228.2146 242.48729 228.25204 242.19107 228.3269 curveto -241.89485 228.40178 241.61001 228.51408 241.33658 228.66382 curveto -241.33658 227.83374 lineto -241.66535 227.70679 241.98436 227.61239 242.29361 227.55054 curveto -242.60285 227.48544 242.90396 227.45289 243.19693 227.45288 curveto -243.98794 227.45289 244.57876 227.65796 244.96939 228.06812 curveto -245.36001 228.47828 245.55532 229.10002 245.55533 229.93335 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -246.76627 227.58472 moveto -247.71841 227.58472 lineto -249.4274 232.17456 lineto -251.13638 227.58472 lineto -252.08853 227.58472 lineto -250.03775 233.05347 lineto -248.81705 233.05347 lineto -246.76627 227.58472 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -258.0065 230.09448 moveto -258.0065 230.53394 lineto -253.87564 230.53394 lineto -253.9147 231.15243 254.10025 231.62443 254.43228 231.94995 curveto -254.76757 232.27222 255.23306 232.43335 255.82877 232.43335 curveto -256.17381 232.43335 256.50747 232.39103 256.82974 232.3064 curveto -257.15526 232.22176 257.47752 232.09481 257.79654 231.92554 curveto -257.79654 232.77515 lineto -257.47427 232.91187 257.14387 233.01603 256.80533 233.08765 curveto -256.46678 233.15926 256.12336 233.19507 255.77505 233.19507 curveto -254.90266 233.19507 254.21093 232.94116 253.69986 232.43335 curveto -253.19205 231.92554 252.93814 231.23869 252.93814 230.3728 curveto -252.93814 229.47762 253.17903 228.76799 253.6608 228.2439 curveto -254.14582 227.71656 254.79849 227.45289 255.6188 227.45288 curveto -256.35448 227.45289 256.93553 227.69052 257.36197 228.16577 curveto -257.79165 228.63778 258.00649 229.28068 258.0065 230.09448 curveto -257.10806 229.83081 moveto -257.10155 229.33928 256.9632 228.94703 256.69302 228.65405 curveto -256.42609 228.36109 256.07128 228.2146 255.62857 228.2146 curveto -255.12727 228.2146 254.72525 228.35621 254.42252 228.6394 curveto -254.12303 228.92261 253.95051 229.32137 253.90494 229.83569 curveto -257.10806 229.83081 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -238.41666 242.74585 moveto -238.41666 243.59546 lineto -238.16275 243.46526 237.89907 243.3676 237.62564 243.30249 curveto -237.3522 243.23739 237.069 243.20484 236.77603 243.20483 curveto -236.33007 243.20484 235.99478 243.2732 235.77017 243.40991 curveto -235.54882 243.54664 235.43814 243.75171 235.43814 244.02515 curveto -235.43814 244.23348 235.51789 244.39787 235.6774 244.51831 curveto -235.8369 244.6355 236.15754 244.74781 236.63931 244.85522 curveto -236.94693 244.92358 lineto -237.58495 245.06031 238.03742 245.25399 238.30435 245.50464 curveto -238.57453 245.75204 238.70962 246.09872 238.70963 246.54468 curveto -238.70962 247.05249 238.5078 247.45451 238.10416 247.75073 curveto -237.70376 248.04696 237.152 248.19507 236.44888 248.19507 curveto -236.15591 248.19507 235.84992 248.16577 235.53091 248.10718 curveto -235.21516 248.05184 234.8815 247.9672 234.52994 247.85327 curveto -234.52994 246.92554 lineto -234.86197 247.09806 235.18912 247.22827 235.51138 247.31616 curveto -235.83365 247.4008 236.15266 247.44312 236.46841 247.44312 curveto -236.89159 247.44312 237.21711 247.3715 237.44498 247.22827 curveto -237.67284 247.08179 237.78677 246.87671 237.78677 246.61304 curveto -237.78677 246.3689 237.70376 246.18172 237.53775 246.05151 curveto -237.37499 245.92131 237.01529 245.79598 236.45865 245.67554 curveto -236.14615 245.60229 lineto -235.58951 245.48511 235.18749 245.30607 234.94009 245.06519 curveto -234.6927 244.82105 234.569 244.48739 234.569 244.06421 curveto -234.569 243.54989 234.75129 243.15276 235.11588 242.8728 curveto -235.48046 242.59286 235.99803 242.45289 236.66861 242.45288 curveto -237.00064 242.45289 237.31314 242.4773 237.60611 242.52612 curveto -237.89907 242.57496 238.16926 242.6482 238.41666 242.74585 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -244.69107 244.75269 moveto -244.69107 248.05347 lineto -243.79263 248.05347 lineto -243.79263 244.78198 lineto -243.79263 244.26441 243.69172 243.87704 243.4899 243.61987 curveto -243.28807 243.36272 242.98534 243.23414 242.5817 243.23413 curveto -242.09667 243.23414 241.71418 243.38876 241.43423 243.698 curveto -241.15428 244.00725 241.01431 244.4288 241.01431 244.96265 curveto -241.01431 248.05347 lineto -240.11099 248.05347 lineto -240.11099 240.45581 lineto -241.01431 240.45581 lineto -241.01431 243.43433 lineto -241.22915 243.10556 241.48143 242.85979 241.77115 242.69702 curveto -242.06411 242.53427 242.40103 242.45289 242.78189 242.45288 curveto -243.41014 242.45289 243.8854 242.6482 244.20767 243.03882 curveto -244.52993 243.42619 244.69107 243.99748 244.69107 244.75269 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -248.61197 243.2146 moveto -248.1302 243.2146 247.74934 243.40341 247.46939 243.78101 curveto -247.18944 244.15536 247.04947 244.66968 247.04947 245.32397 curveto -247.04947 245.97827 247.18781 246.49422 247.46451 246.87183 curveto -247.74445 247.24618 248.12694 247.43335 248.61197 247.43335 curveto -249.09048 247.43335 249.46971 247.24455 249.74966 246.86694 curveto -250.02961 246.48934 250.16958 245.97502 250.16959 245.32397 curveto -250.16958 244.67619 250.02961 244.1635 249.74966 243.78589 curveto -249.46971 243.40503 249.09048 243.2146 248.61197 243.2146 curveto -248.61197 242.45288 moveto -249.39322 242.45289 250.00682 242.70679 250.45279 243.2146 curveto -250.89875 243.72242 251.12173 244.42554 251.12173 245.32397 curveto -251.12173 246.21916 250.89875 246.92228 250.45279 247.43335 curveto -250.00682 247.94116 249.39322 248.19507 248.61197 248.19507 curveto -247.82746 248.19507 247.21223 247.94116 246.76627 247.43335 curveto -246.32356 246.92228 246.1022 246.21916 246.1022 245.32397 curveto -246.1022 244.42554 246.32356 243.72242 246.76627 243.2146 curveto -247.21223 242.70679 247.82746 242.45289 248.61197 242.45288 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -252.08365 242.58472 moveto -252.98209 242.58472 lineto -254.10513 246.85229 lineto -255.2233 242.58472 lineto -256.28287 242.58472 lineto -257.40591 246.85229 lineto -258.52408 242.58472 lineto -259.42252 242.58472 lineto -257.99185 248.05347 lineto -256.93228 248.05347 lineto -255.75552 243.57104 lineto -254.57388 248.05347 lineto -253.51431 248.05347 lineto -252.08365 242.58472 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -311.38464 185.46135 moveto -311.38464 188.76213 lineto -310.48621 188.76213 lineto -310.48621 185.49065 lineto -310.4862 184.97307 310.38529 184.5857 310.18347 184.32854 curveto -309.98164 184.07138 309.67891 183.9428 309.27527 183.94279 curveto -308.79024 183.9428 308.40775 184.09742 308.12781 184.40666 curveto -307.84786 184.71591 307.70788 185.13746 307.70789 185.67131 curveto -307.70789 188.76213 lineto -306.80457 188.76213 lineto -306.80457 181.16447 lineto -307.70789 181.16447 lineto -307.70789 184.14299 lineto -307.92273 183.81422 308.17501 183.56845 308.46472 183.40569 curveto -308.75769 183.24293 309.0946 183.16155 309.47546 183.16154 curveto -310.10371 183.16155 310.57897 183.35686 310.90125 183.74748 curveto -311.22351 184.13486 311.38464 184.70615 311.38464 185.46135 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -315.67175 186.01311 moveto -314.94584 186.01311 314.44291 186.09612 314.16296 186.26213 curveto -313.88301 186.42815 313.74304 186.71135 313.74304 187.11174 curveto -313.74304 187.43075 313.84721 187.68466 314.05554 187.87346 curveto -314.26713 188.05901 314.55359 188.15178 314.91492 188.15178 curveto -315.41296 188.15178 315.81172 187.976 316.11121 187.62444 curveto -316.41394 187.26962 316.5653 186.79924 316.56531 186.2133 curveto -316.56531 186.01311 lineto -315.67175 186.01311 lineto -317.46375 185.64201 moveto -317.46375 188.76213 lineto -316.56531 188.76213 lineto -316.56531 187.93205 lineto -316.36023 188.26408 316.10469 188.50985 315.79871 188.66936 curveto -315.49271 188.82561 315.11836 188.90373 314.67566 188.90373 curveto -314.11576 188.90373 313.6698 188.74748 313.33777 188.43498 curveto -313.00899 188.11923 312.8446 187.69768 312.8446 187.17033 curveto -312.8446 186.5551 313.04968 186.09123 313.45984 185.77873 curveto -313.87325 185.46624 314.48848 185.30999 315.30554 185.30998 curveto -316.56531 185.30998 lineto -316.56531 185.22209 lineto -316.5653 184.80868 316.42858 184.48967 316.15515 184.26506 curveto -315.88497 184.0372 315.50411 183.92327 315.01257 183.92326 curveto -314.70007 183.92327 314.39571 183.9607 314.09949 184.03557 curveto -313.80326 184.11044 313.51843 184.22275 313.245 184.37248 curveto -313.245 183.5424 lineto -313.57377 183.41546 313.89278 183.32106 314.20203 183.2592 curveto -314.51127 183.1941 314.81238 183.16155 315.10535 183.16154 curveto -315.89636 183.16155 316.48718 183.36663 316.87781 183.77678 curveto -317.26843 184.18694 317.46374 184.80868 317.46375 185.64201 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -318.67468 183.29338 moveto -319.62683 183.29338 lineto -321.33582 187.88322 lineto -323.0448 183.29338 lineto -323.99695 183.29338 lineto -321.94617 188.76213 lineto -320.72546 188.76213 lineto -318.67468 183.29338 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -329.91492 185.80315 moveto -329.91492 186.2426 lineto -325.78406 186.2426 lineto -325.82312 186.86109 326.00867 187.3331 326.3407 187.65862 curveto -326.67598 187.98088 327.14148 188.14201 327.73718 188.14201 curveto -328.08223 188.14201 328.41589 188.0997 328.73816 188.01506 curveto -329.06368 187.93043 329.38594 187.80347 329.70496 187.6342 curveto -329.70496 188.48381 lineto -329.38269 188.62053 329.05228 188.7247 328.71375 188.79631 curveto -328.3752 188.86792 328.03178 188.90373 327.68347 188.90373 curveto -326.81107 188.90373 326.11934 188.64983 325.60828 188.14201 curveto -325.10046 187.6342 324.84656 186.94735 324.84656 186.08147 curveto -324.84656 185.18629 325.08744 184.47665 325.56921 183.95256 curveto -326.05424 183.42522 326.70691 183.16155 327.52722 183.16154 curveto -328.26289 183.16155 328.84395 183.39918 329.27039 183.87444 curveto -329.70007 184.34645 329.91491 184.98935 329.91492 185.80315 curveto -329.01648 185.53947 moveto -329.00996 185.04794 328.87162 184.65569 328.60144 184.36272 curveto -328.33451 184.06975 327.97969 183.92327 327.53699 183.92326 curveto -327.03568 183.92327 326.63366 184.06487 326.33093 184.34807 curveto -326.03145 184.63128 325.85893 185.03004 325.81335 185.54436 curveto -329.01648 185.53947 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -310.32507 198.45451 moveto -310.32507 199.30412 lineto -310.07116 199.17392 309.80749 199.07626 309.53406 199.01115 curveto -309.26062 198.94605 308.97741 198.9135 308.68445 198.9135 curveto -308.23848 198.9135 307.9032 198.98186 307.67859 199.11858 curveto -307.45723 199.2553 307.34656 199.46038 307.34656 199.73381 curveto -307.34656 199.94215 307.42631 200.10654 307.58582 200.22697 curveto -307.74532 200.34417 308.06596 200.45647 308.54773 200.56389 curveto -308.85535 200.63225 lineto -309.49336 200.76897 309.94584 200.96265 310.21277 201.2133 curveto -310.48295 201.4607 310.61804 201.80738 310.61804 202.25334 curveto -310.61804 202.76116 310.41621 203.16317 310.01257 203.4594 curveto -309.61218 203.75562 309.06042 203.90373 308.3573 203.90373 curveto -308.06433 203.90373 307.75834 203.87444 307.43933 203.81584 curveto -307.12357 203.7605 306.78992 203.67587 306.43835 203.56194 curveto -306.43835 202.6342 lineto -306.77038 202.80673 307.09753 202.93694 307.4198 203.02483 curveto -307.74206 203.10946 308.06107 203.15178 308.37683 203.15178 curveto -308.80001 203.15178 309.12553 203.08017 309.35339 202.93694 curveto -309.58125 202.79045 309.69519 202.58537 309.69519 202.3217 curveto -309.69519 202.07756 309.61218 201.89039 309.44617 201.76018 curveto -309.2834 201.62997 308.9237 201.50465 308.36707 201.3842 curveto -308.05457 201.31096 lineto -307.49792 201.19377 307.09591 201.01474 306.84851 200.77385 curveto -306.60111 200.52971 306.47742 200.19605 306.47742 199.77287 curveto -306.47742 199.25855 306.65971 198.86142 307.02429 198.58147 curveto -307.38887 198.30152 307.90645 198.16155 308.57703 198.16154 curveto -308.90905 198.16155 309.22155 198.18596 309.51453 198.23479 curveto -309.80749 198.28362 310.07767 198.35686 310.32507 198.45451 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -316.59949 200.46135 moveto -316.59949 203.76213 lineto -315.70105 203.76213 lineto -315.70105 200.49065 lineto -315.70105 199.97307 315.60013 199.5857 315.39832 199.32854 curveto -315.19649 199.07138 314.89375 198.9428 314.49011 198.94279 curveto -314.00508 198.9428 313.6226 199.09742 313.34265 199.40666 curveto -313.0627 199.71591 312.92273 200.13746 312.92273 200.67131 curveto -312.92273 203.76213 lineto -312.01941 203.76213 lineto -312.01941 196.16447 lineto -312.92273 196.16447 lineto -312.92273 199.14299 lineto -313.13757 198.81422 313.38985 198.56845 313.67957 198.40569 curveto -313.97253 198.24293 314.30945 198.16155 314.69031 198.16154 curveto -315.31856 198.16155 315.79382 198.35686 316.11609 198.74748 curveto -316.43835 199.13486 316.59948 199.70615 316.59949 200.46135 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -320.52039 198.92326 moveto -320.03861 198.92327 319.65775 199.11207 319.37781 199.48967 curveto -319.09786 199.86402 318.95788 200.37835 318.95789 201.03264 curveto -318.95788 201.68694 319.09623 202.20289 319.37292 202.58049 curveto -319.65287 202.95484 320.03536 203.14201 320.52039 203.14201 curveto -320.9989 203.14201 321.37813 202.95321 321.65808 202.57561 curveto -321.93802 202.198 322.078 201.68368 322.078 201.03264 curveto -322.078 200.38486 321.93802 199.87216 321.65808 199.49455 curveto -321.37813 199.1137 320.9989 198.92327 320.52039 198.92326 curveto -320.52039 198.16154 moveto -321.30163 198.16155 321.91524 198.41546 322.36121 198.92326 curveto -322.80716 199.43108 323.03015 200.1342 323.03015 201.03264 curveto -323.03015 201.92782 322.80716 202.63095 322.36121 203.14201 curveto -321.91524 203.64983 321.30163 203.90373 320.52039 203.90373 curveto -319.73588 203.90373 319.12064 203.64983 318.67468 203.14201 curveto -318.23197 202.63095 318.01062 201.92782 318.01062 201.03264 curveto -318.01062 200.1342 318.23197 199.43108 318.67468 198.92326 curveto -319.12064 198.41546 319.73588 198.16155 320.52039 198.16154 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -323.99207 198.29338 moveto -324.8905 198.29338 lineto -326.01355 202.56096 lineto -327.13171 198.29338 lineto -328.19128 198.29338 lineto -329.31433 202.56096 lineto -330.4325 198.29338 lineto -331.33093 198.29338 lineto -329.90027 203.76213 lineto -328.8407 203.76213 lineto -327.66394 199.27971 lineto -326.4823 203.76213 lineto -325.42273 203.76213 lineto -323.99207 198.29338 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -305.63477 140.25864 moveto -305.63477 143.15903 lineto -304.73145 143.15903 lineto -304.73145 135.6102 lineto -305.63477 135.6102 lineto -305.63477 136.44028 lineto -305.82357 136.11476 306.0612 135.87388 306.34766 135.71762 curveto -306.63737 135.55812 306.98242 135.47837 307.38281 135.47836 curveto -308.04687 135.47837 308.58561 135.74204 308.99902 136.26938 curveto -309.41568 136.79673 309.62402 137.49009 309.62402 138.34946 curveto -309.62402 139.20883 309.41568 139.90219 308.99902 140.42953 curveto -308.58561 140.95688 308.04687 141.22055 307.38281 141.22055 curveto -306.98242 141.22055 306.63737 141.14243 306.34766 140.98618 curveto -306.0612 140.82667 305.82357 140.58416 305.63477 140.25864 curveto -308.69141 138.34946 moveto -308.6914 137.68865 308.55468 137.17108 308.28125 136.79672 curveto -308.01106 136.41912 307.63834 136.23032 307.16309 136.23032 curveto -306.68782 136.23032 306.31347 136.41912 306.04004 136.79672 curveto -305.76985 137.17108 305.63476 137.68865 305.63477 138.34946 curveto -305.63476 139.01027 305.76985 139.52947 306.04004 139.90707 curveto -306.31347 140.28142 306.68782 140.4686 307.16309 140.4686 curveto -307.63834 140.4686 308.01106 140.28142 308.28125 139.90707 curveto -308.55468 139.52947 308.6914 139.01027 308.69141 138.34946 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -314.28223 136.45004 moveto -314.18131 136.39145 314.07063 136.34914 313.9502 136.32309 curveto -313.833 136.2938 313.7028 136.27915 313.55957 136.27914 curveto -313.05175 136.27915 312.66113 136.44516 312.3877 136.77719 curveto -312.11751 137.10597 311.98242 137.5796 311.98242 138.19809 curveto -311.98242 141.07895 lineto -311.0791 141.07895 lineto -311.0791 135.6102 lineto -311.98242 135.6102 lineto -311.98242 136.45981 lineto -312.17122 136.12778 312.41699 135.88201 312.71973 135.7225 curveto -313.02246 135.55975 313.3903 135.47837 313.82324 135.47836 curveto -313.88509 135.47837 313.95345 135.48325 314.02832 135.49301 curveto -314.10319 135.49953 314.18619 135.51092 314.27734 135.52719 curveto -314.28223 136.45004 lineto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -317.13867 136.24008 moveto -316.6569 136.24009 316.27604 136.42889 315.99609 136.80649 curveto -315.71614 137.18084 315.57617 137.69516 315.57617 138.34946 curveto -315.57617 139.00376 315.71452 139.51971 315.99121 139.89731 curveto -316.27116 140.27166 316.65364 140.45883 317.13867 140.45883 curveto -317.61718 140.45883 317.99642 140.27003 318.27637 139.89243 curveto -318.55631 139.51482 318.69628 139.0005 318.69629 138.34946 curveto -318.69628 137.70167 318.55631 137.18898 318.27637 136.81137 curveto -317.99642 136.43052 317.61718 136.24009 317.13867 136.24008 curveto -317.13867 135.47836 moveto -317.91992 135.47837 318.53352 135.73227 318.97949 136.24008 curveto -319.42545 136.7479 319.64843 137.45102 319.64844 138.34946 curveto -319.64843 139.24464 319.42545 139.94777 318.97949 140.45883 curveto -318.53352 140.96664 317.91992 141.22055 317.13867 141.22055 curveto -316.35416 141.22055 315.73893 140.96664 315.29297 140.45883 curveto -314.85026 139.94777 314.62891 139.24464 314.62891 138.34946 curveto -314.62891 137.45102 314.85026 136.7479 315.29297 136.24008 curveto -315.73893 135.73227 316.35416 135.47837 317.13867 135.47836 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -323.25195 136.24008 moveto -322.77018 136.24009 322.38932 136.42889 322.10938 136.80649 curveto -321.82943 137.18084 321.68945 137.69516 321.68945 138.34946 curveto -321.68945 139.00376 321.8278 139.51971 322.10449 139.89731 curveto -322.38444 140.27166 322.76692 140.45883 323.25195 140.45883 curveto -323.73047 140.45883 324.1097 140.27003 324.38965 139.89243 curveto -324.66959 139.51482 324.80957 139.0005 324.80957 138.34946 curveto -324.80957 137.70167 324.66959 137.18898 324.38965 136.81137 curveto -324.1097 136.43052 323.73047 136.24009 323.25195 136.24008 curveto -323.25195 135.47836 moveto -324.0332 135.47837 324.64681 135.73227 325.09277 136.24008 curveto -325.53873 136.7479 325.76171 137.45102 325.76172 138.34946 curveto -325.76171 139.24464 325.53873 139.94777 325.09277 140.45883 curveto -324.64681 140.96664 324.0332 141.22055 323.25195 141.22055 curveto -322.46745 141.22055 321.85221 140.96664 321.40625 140.45883 curveto -320.96354 139.94777 320.74219 139.24464 320.74219 138.34946 curveto -320.74219 137.45102 320.96354 136.7479 321.40625 136.24008 curveto -321.85221 135.73227 322.46745 135.47837 323.25195 135.47836 curveto -fill -grestore -gsave -0 0 0 setrgbcolor -newpath -330.01465 133.48129 moveto -330.01465 134.22836 lineto -329.15527 134.22836 lineto -328.83301 134.22837 328.6084 134.29347 328.48145 134.42368 curveto -328.35775 134.55389 328.2959 134.78827 328.2959 135.1268 curveto -328.2959 135.6102 lineto -329.77539 135.6102 lineto -329.77539 136.30844 lineto -328.2959 136.30844 lineto -328.2959 141.07895 lineto -327.39258 141.07895 lineto -327.39258 136.30844 lineto -326.5332 136.30844 lineto -326.5332 135.6102 lineto -327.39258 135.6102 lineto -327.39258 135.22934 lineto -327.39258 134.62062 327.53418 134.17791 327.81738 133.90121 curveto -328.10058 133.62127 328.5498 133.4813 329.16504 133.48129 curveto -330.01465 133.48129 lineto -fill -grestore -grestore -grestore -showpage -%%EOF diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/isar-vm.pdf Binary file doc-src/IsarRef/Thy/document/isar-vm.pdf has changed diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/Thy/document/isar-vm.svg --- a/doc-src/IsarRef/Thy/document/isar-vm.svg Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,460 +0,0 @@ - - - - - - - - - - - - - - - - - - - - - - - - - image/svg+xml - - - - - - - - chain - - - - - - - - prove - - - - - - state - - - - - - - - theorem - - qed - qed - fixassume - { }next - notelet - usingunfolding - then - haveshow - haveshow - proof - - diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/isar-vm.eps --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/document/isar-vm.eps Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,2694 @@ +%!PS-Adobe-3.0 EPSF-3.0 +%%Creator: inkscape 0.46 +%%Pages: 1 +%%Orientation: Portrait +%%BoundingBox: 0 0 435 173 +%%HiResBoundingBox: 0 0 435 173 +%%EndComments +%%BeginSetup +%%EndSetup +%%Page: 1 1 +0 173 translate +0.8 -0.8 scale +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +gsave [1 0 0 1 0 0] concat +gsave [1 0 0 1 -44.641342 -76.87234] concat +gsave [1 0 0 1 70.838012 79.725562] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +229.77649 131.52507 moveto +265.28729 131.52507 lineto +275.08072 131.52507 282.96496 139.40931 282.96496 149.20274 curveto +282.96496 166.99701 lineto +282.96496 176.79043 275.08072 184.67467 265.28729 184.67467 curveto +229.77649 184.67467 lineto +219.98306 184.67467 212.09882 176.79043 212.09882 166.99701 curveto +212.09882 149.20274 lineto +212.09882 139.40931 219.98306 131.52507 229.77649 131.52507 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +231.92252 155.58815 moveto +231.92252 157.8694 lineto +231.5423 157.60899 231.15949 157.41628 230.77408 157.29128 curveto +230.39386 157.16628 229.99803 157.10378 229.58658 157.10378 curveto +228.80532 157.10378 228.19595 157.33295 227.75845 157.79128 curveto +227.32616 158.24441 227.11001 158.87982 227.11002 159.69753 curveto +227.11001 160.51524 227.32616 161.15326 227.75845 161.61159 curveto +228.19595 162.06471 228.80532 162.29128 229.58658 162.29128 curveto +230.02407 162.29128 230.43813 162.22617 230.82877 162.09596 curveto +231.22459 161.96576 231.58917 161.77305 231.92252 161.51784 curveto +231.92252 163.8069 lineto +231.48501 163.96836 231.0397 164.08815 230.58658 164.16628 curveto +230.13866 164.24961 229.68813 164.29127 229.23502 164.29128 curveto +227.65689 164.29127 226.42251 163.88763 225.53189 163.08034 curveto +224.64126 162.26784 224.19595 161.14024 224.19595 159.69753 curveto +224.19595 158.25482 224.64126 157.12982 225.53189 156.32253 curveto +226.42251 155.51003 227.65689 155.10378 229.23502 155.10378 curveto +229.69334 155.10378 230.14386 155.14545 230.58658 155.22878 curveto +231.03449 155.30691 231.4798 155.4267 231.92252 155.58815 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +243.14908 158.73659 moveto +243.14908 164.06471 lineto +240.33658 164.06471 lineto +240.33658 163.19753 lineto +240.33658 160.00221 lineto +240.33657 159.23659 240.31834 158.71055 240.28189 158.42409 curveto +240.25063 158.13764 240.19334 157.9267 240.11002 157.79128 curveto +240.00063 157.60899 239.8522 157.46836 239.6647 157.3694 curveto +239.4772 157.26524 239.26366 157.21316 239.02408 157.21315 curveto +238.44074 157.21316 237.98241 157.43972 237.64908 157.89284 curveto +237.31574 158.34076 237.14907 158.96316 237.14908 159.76003 curveto +237.14908 164.06471 lineto +234.3522 164.06471 lineto +234.3522 151.90846 lineto +237.14908 151.90846 lineto +237.14908 156.59596 lineto +237.57095 156.08555 238.01887 155.71055 238.49283 155.47096 curveto +238.96678 155.22618 239.49022 155.10378 240.06314 155.10378 curveto +241.07355 155.10378 241.83917 155.41368 242.36002 156.03346 curveto +242.88605 156.65326 243.14907 157.5543 243.14908 158.73659 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +249.68033 160.12721 moveto +249.09699 160.12722 248.65689 160.22617 248.36002 160.42409 curveto +248.06835 160.62201 247.92251 160.91367 247.92252 161.29909 curveto +247.92251 161.65326 248.0397 161.9319 248.27408 162.13503 curveto +248.51366 162.33294 248.84439 162.4319 249.26627 162.4319 curveto +249.7923 162.4319 250.23501 162.2444 250.59439 161.8694 curveto +250.95376 161.48919 251.13345 161.01524 251.13345 160.44753 curveto +251.13345 160.12721 lineto +249.68033 160.12721 lineto +253.95377 159.07253 moveto +253.95377 164.06471 lineto +251.13345 164.06471 lineto +251.13345 162.76784 lineto +250.75845 163.29909 250.33657 163.68711 249.86783 163.9319 curveto +249.39907 164.17148 248.82876 164.29127 248.15689 164.29128 curveto +247.25064 164.29127 246.51366 164.02825 245.94595 163.50221 curveto +245.38345 162.97096 245.1022 162.28346 245.1022 161.43971 curveto +245.1022 160.41367 245.45376 159.66107 246.15689 159.1819 curveto +246.86522 158.70274 247.9746 158.46316 249.48502 158.46315 curveto +251.13345 158.46315 lineto +251.13345 158.2444 lineto +251.13345 157.8017 250.95897 157.47878 250.61002 157.27565 curveto +250.26105 157.06732 249.71678 156.96316 248.9772 156.96315 curveto +248.37824 156.96316 247.82095 157.02305 247.30533 157.14284 curveto +246.7897 157.26264 246.31053 157.44232 245.86783 157.6819 curveto +245.86783 155.54909 lineto +246.46678 155.40326 247.06835 155.29389 247.67252 155.22096 curveto +248.27668 155.14285 248.88084 155.10378 249.48502 155.10378 curveto +251.06313 155.10378 252.20115 155.41628 252.89908 156.04128 curveto +253.60219 156.66107 253.95376 157.67149 253.95377 159.07253 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +256.57095 155.31471 moveto +259.36783 155.31471 lineto +259.36783 164.06471 lineto +256.57095 164.06471 lineto +256.57095 155.31471 lineto +256.57095 151.90846 moveto +259.36783 151.90846 lineto +259.36783 154.18971 lineto +256.57095 154.18971 lineto +256.57095 151.90846 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +270.86783 158.73659 moveto +270.86783 164.06471 lineto +268.05533 164.06471 lineto +268.05533 163.19753 lineto +268.05533 159.98659 lineto +268.05532 159.23138 268.03709 158.71055 268.00064 158.42409 curveto +267.96938 158.13764 267.91209 157.9267 267.82877 157.79128 curveto +267.71938 157.60899 267.57095 157.46836 267.38345 157.3694 curveto +267.19595 157.26524 266.98241 157.21316 266.74283 157.21315 curveto +266.15949 157.21316 265.70116 157.43972 265.36783 157.89284 curveto +265.03449 158.34076 264.86782 158.96316 264.86783 159.76003 curveto +264.86783 164.06471 lineto +262.07095 164.06471 lineto +262.07095 155.31471 lineto +264.86783 155.31471 lineto +264.86783 156.59596 lineto +265.2897 156.08555 265.73762 155.71055 266.21158 155.47096 curveto +266.68553 155.22618 267.20897 155.10378 267.78189 155.10378 curveto +268.7923 155.10378 269.55792 155.41368 270.07877 156.03346 curveto +270.6048 156.65326 270.86782 157.5543 270.86783 158.73659 curveto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +0 setlinejoin +0 setlinecap +newpath +424.72469 236.82544 moveto +356.83209 236.82544 lineto +356.83209 236.82544 lineto +stroke +gsave [-0.39968505 4.8945685e-17 -4.8945685e-17 -0.39968505 356.83209 236.82544] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99921268 setlinewidth +0 setlinejoin +0 setlinecap +newpath +282.35183 236.82544 moveto +215.11403 236.82544 lineto +215.11403 236.82544 lineto +stroke +gsave [-0.39968507 4.8945688e-17 -4.8945688e-17 -0.39968507 215.11403 236.82544] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +0.99999994 setlinewidth +0 setlinejoin +0 setlinecap +newpath +424.69726 192.5341 moveto +215.13005 192.5341 lineto +stroke +gsave [-0.39999998 4.8984251e-17 -4.8984251e-17 -0.39999998 215.13005 192.5341] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +211.98429 148.24276 moveto +422.13162 148.24276 lineto +stroke +gsave [0.4 0 0 0.4 422.13162 148.24276] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [1 0 0 1 70.866146 78.725567] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +88.044201 42.942394 moveto +123.555 42.942394 lineto +133.34843 42.942394 141.23267 50.826635 141.23267 60.620064 curveto +141.23267 166.99701 lineto +141.23267 176.79044 133.34843 184.67468 123.555 184.67468 curveto +88.044201 184.67468 lineto +78.250772 184.67468 70.366531 176.79044 70.366531 166.99701 curveto +70.366531 60.620064 lineto +70.366531 50.826635 78.250772 42.942394 88.044201 42.942394 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +83.823044 115.35931 moveto +83.823044 119.95306 lineto +81.026169 119.95306 lineto +81.026169 107.87494 lineto +83.823044 107.87494 lineto +83.823044 109.15619 lineto +84.208456 108.64578 84.635539 108.27078 85.104294 108.03119 curveto +85.573038 107.78641 86.1121 107.66401 86.721481 107.664 curveto +87.799598 107.66401 88.685014 108.0937 89.377731 108.95306 curveto +90.070429 109.80724 90.416783 110.9088 90.416794 112.25775 curveto +90.416783 113.60671 90.070429 114.71088 89.377731 115.57025 curveto +88.685014 116.42442 87.799598 116.8515 86.721481 116.8515 curveto +86.1121 116.8515 85.573038 116.73171 85.104294 116.49213 curveto +84.635539 116.24734 84.208456 115.86973 83.823044 115.35931 curveto +85.682419 109.69525 moveto +85.083455 109.69526 84.622518 109.91661 84.299606 110.35931 curveto +83.981894 110.79682 83.82304 111.42963 83.823044 112.25775 curveto +83.82304 113.08588 83.981894 113.7213 84.299606 114.164 curveto +84.622518 114.6015 85.083455 114.82025 85.682419 114.82025 curveto +86.281371 114.82025 86.737099 114.6015 87.049606 114.164 curveto +87.367307 113.7265 87.526161 113.09109 87.526169 112.25775 curveto +87.526161 111.42442 87.367307 110.78901 87.049606 110.3515 curveto +86.737099 109.91401 86.281371 109.69526 85.682419 109.69525 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +98.994919 110.25775 moveto +98.75012 110.14317 98.505328 110.05984 98.260544 110.00775 curveto +98.020954 109.95047 97.778766 109.92182 97.533981 109.92181 curveto +96.815226 109.92182 96.260539 110.15359 95.869919 110.61713 curveto +95.484498 111.07547 95.29179 111.73432 95.291794 112.59369 curveto +95.291794 116.62494 lineto +92.494919 116.62494 lineto +92.494919 107.87494 lineto +95.291794 107.87494 lineto +95.291794 109.31244 lineto +95.651164 108.73953 96.062622 108.32286 96.526169 108.06244 curveto +96.994913 107.79682 97.554808 107.66401 98.205856 107.664 curveto +98.299599 107.66401 98.401162 107.66922 98.510544 107.67963 curveto +98.619911 107.68484 98.778765 107.70047 98.987106 107.7265 curveto +98.994919 110.25775 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +104.56523 109.664 moveto +103.94543 109.66401 103.47148 109.88797 103.14336 110.33588 curveto +102.82044 110.77859 102.65898 111.41922 102.65898 112.25775 curveto +102.65898 113.0963 102.82044 113.73953 103.14336 114.18744 curveto +103.47148 114.63015 103.94543 114.8515 104.56523 114.8515 curveto +105.1746 114.8515 105.64075 114.63015 105.96367 114.18744 curveto +106.28658 113.73953 106.44804 113.0963 106.44804 112.25775 curveto +106.44804 111.41922 106.28658 110.77859 105.96367 110.33588 curveto +105.64075 109.88797 105.1746 109.66401 104.56523 109.664 curveto +104.56523 107.664 moveto +106.07043 107.66401 107.24491 108.07026 108.08867 108.88275 curveto +108.93762 109.69526 109.3621 110.82026 109.36211 112.25775 curveto +109.3621 113.69525 108.93762 114.82025 108.08867 115.63275 curveto +107.24491 116.44525 106.07043 116.8515 104.56523 116.8515 curveto +103.05481 116.8515 101.87252 116.44525 101.01836 115.63275 curveto +100.1694 114.82025 99.744918 113.69525 99.744919 112.25775 curveto +99.744918 110.82026 100.1694 109.69526 101.01836 108.88275 curveto +101.87252 108.07026 103.05481 107.66401 104.56523 107.664 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +110.29961 107.87494 moveto +113.09648 107.87494 lineto +115.27617 113.92181 lineto +117.44804 107.87494 lineto +120.25273 107.87494 lineto +116.80742 116.62494 lineto +113.73711 116.62494 lineto +110.29961 107.87494 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +130.57304 112.2265 moveto +130.57304 113.02338 lineto +124.03398 113.02338 lineto +124.10169 113.67963 124.33866 114.17182 124.74492 114.49994 curveto +125.15116 114.82807 125.71887 114.99213 126.44804 114.99213 curveto +127.03658 114.99213 127.63814 114.90619 128.25273 114.73431 curveto +128.87251 114.55723 129.50793 114.29161 130.15898 113.93744 curveto +130.15898 116.09369 lineto +129.49751 116.34369 128.83606 116.53119 128.17461 116.65619 curveto +127.51314 116.7864 126.85168 116.8515 126.19023 116.8515 curveto +124.60689 116.8515 123.37512 116.45046 122.49492 115.64838 curveto +121.61992 114.84109 121.18242 113.71088 121.18242 112.25775 curveto +121.18242 110.83067 121.61211 109.70828 122.47148 108.89056 curveto +123.33606 108.07286 124.52356 107.66401 126.03398 107.664 curveto +127.40897 107.66401 128.50793 108.07807 129.33086 108.90619 curveto +130.15897 109.73432 130.57303 110.84109 130.57304 112.2265 curveto +127.69804 111.29681 moveto +127.69804 110.76557 127.54179 110.33849 127.22929 110.01556 curveto +126.922 109.68745 126.51835 109.52338 126.01836 109.52338 curveto +125.47668 109.52338 125.03658 109.67703 124.69804 109.98431 curveto +124.3595 110.2864 124.14856 110.7239 124.06523 111.29681 curveto +127.69804 111.29681 lineto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +176.66575 92.035445 moveto +176.66575 118.61025 lineto +stroke +gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 176.66575 118.61025] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [0.2378166 0 0 -0.2269133 90.621413 253.06251] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [1 0 0 1 70.866151 78.725565] concat +0 0 0 setrgbcolor +[] 0 setdash +0.99921262 setlinewidth +1 setlinejoin +1 setlinecap +newpath +371.50879 42.942394 moveto +407.01959 42.942394 lineto +416.81302 42.942394 424.69726 50.826635 424.69726 60.620064 curveto +424.69726 166.99701 lineto +424.69726 176.79044 416.81302 184.67468 407.01959 184.67468 curveto +371.50879 184.67468 lineto +361.71536 184.67468 353.83112 176.79044 353.83112 166.99701 curveto +353.83112 60.620064 lineto +353.83112 50.826635 361.71536 42.942394 371.50879 42.942394 curveto +closepath +stroke +gsave +0 0 0 setrgbcolor +newpath +374.16263 110.83588 moveto +374.16263 112.96088 lineto +373.56366 112.71088 372.98554 112.52338 372.42825 112.39838 curveto +371.87096 112.27338 371.34491 112.21088 370.85013 112.21088 curveto +370.31887 112.21088 369.92304 112.27859 369.66263 112.414 curveto +369.40742 112.54422 369.27981 112.74734 369.27982 113.02338 curveto +369.27981 113.24734 369.37617 113.41922 369.56888 113.539 curveto +369.76679 113.6588 370.11835 113.74734 370.62357 113.80463 curveto +371.11575 113.87494 lineto +372.54804 114.05724 373.51158 114.35671 374.00638 114.77338 curveto +374.50116 115.19005 374.74856 115.84369 374.74857 116.73431 curveto +374.74856 117.66661 374.40481 118.36713 373.71732 118.83588 curveto +373.02981 119.30463 372.00377 119.539 370.63919 119.539 curveto +370.06106 119.539 369.4621 119.49213 368.84232 119.39838 curveto +368.22773 119.30983 367.59492 119.17442 366.94388 118.99213 curveto +366.94388 116.86713 lineto +367.50117 117.13796 368.07148 117.34109 368.65482 117.4765 curveto +369.24335 117.61192 369.83971 117.67963 370.44388 117.67963 curveto +370.99075 117.67963 371.40221 117.60411 371.67825 117.45306 curveto +371.95429 117.30202 372.09231 117.07807 372.09232 116.78119 curveto +372.09231 116.53119 371.99596 116.3463 371.80325 116.2265 curveto +371.61575 116.1015 371.23814 116.00515 370.67044 115.93744 curveto +370.17825 115.87494 lineto +368.93346 115.71869 368.06106 115.42963 367.56107 115.00775 curveto +367.06106 114.58588 366.81106 113.94526 366.81107 113.08588 curveto +366.81106 112.1588 367.12877 111.4713 367.76419 111.02338 curveto +368.3996 110.57547 369.37356 110.35151 370.68607 110.3515 curveto +371.20169 110.35151 371.74335 110.39057 372.31107 110.46869 curveto +372.87877 110.54682 373.49595 110.66922 374.16263 110.83588 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +379.91263 108.07806 moveto +379.91263 110.56244 lineto +382.79544 110.56244 lineto +382.79544 112.56244 lineto +379.91263 112.56244 lineto +379.91263 116.27338 lineto +379.91262 116.67963 379.99335 116.95567 380.15482 117.1015 curveto +380.31627 117.24213 380.63658 117.31244 381.11575 117.31244 curveto +382.55325 117.31244 lineto +382.55325 119.31244 lineto +380.15482 119.31244 lineto +379.05065 119.31244 378.26679 119.08327 377.80325 118.62494 curveto +377.34492 118.1614 377.11575 117.37755 377.11575 116.27338 curveto +377.11575 112.56244 lineto +375.72513 112.56244 lineto +375.72513 110.56244 lineto +377.11575 110.56244 lineto +377.11575 108.07806 lineto +379.91263 108.07806 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +388.43607 115.37494 moveto +387.85273 115.37494 387.41262 115.4739 387.11575 115.67181 curveto +386.82408 115.86973 386.67825 116.1614 386.67825 116.54681 curveto +386.67825 116.90098 386.79544 117.17963 387.02982 117.38275 curveto +387.26939 117.58067 387.60012 117.67963 388.022 117.67963 curveto +388.54804 117.67963 388.99075 117.49213 389.35013 117.11713 curveto +389.7095 116.73692 389.88918 116.26296 389.88919 115.69525 curveto +389.88919 115.37494 lineto +388.43607 115.37494 lineto +392.7095 114.32025 moveto +392.7095 119.31244 lineto +389.88919 119.31244 lineto +389.88919 118.01556 lineto +389.51418 118.54681 389.09231 118.93484 388.62357 119.17963 curveto +388.15481 119.41921 387.5845 119.539 386.91263 119.539 curveto +386.00638 119.539 385.2694 119.27598 384.70169 118.74994 curveto +384.13919 118.21869 383.85794 117.53119 383.85794 116.68744 curveto +383.85794 115.6614 384.2095 114.9088 384.91263 114.42963 curveto +385.62096 113.95047 386.73033 113.71088 388.24075 113.71088 curveto +389.88919 113.71088 lineto +389.88919 113.49213 lineto +389.88918 113.04942 389.7147 112.72651 389.36575 112.52338 curveto +389.01679 112.31505 388.47252 112.21088 387.73294 112.21088 curveto +387.13398 112.21088 386.57669 112.27078 386.06107 112.39056 curveto +385.54544 112.51036 385.06627 112.69005 384.62357 112.92963 curveto +384.62357 110.79681 lineto +385.22252 110.65099 385.82408 110.54161 386.42825 110.46869 curveto +387.03242 110.39057 387.63658 110.35151 388.24075 110.3515 curveto +389.81887 110.35151 390.95689 110.66401 391.65482 111.289 curveto +392.35793 111.9088 392.70949 112.91922 392.7095 114.32025 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +398.38138 108.07806 moveto +398.38138 110.56244 lineto +401.26419 110.56244 lineto +401.26419 112.56244 lineto +398.38138 112.56244 lineto +398.38138 116.27338 lineto +398.38137 116.67963 398.4621 116.95567 398.62357 117.1015 curveto +398.78502 117.24213 399.10533 117.31244 399.5845 117.31244 curveto +401.022 117.31244 lineto +401.022 119.31244 lineto +398.62357 119.31244 lineto +397.5194 119.31244 396.73554 119.08327 396.272 118.62494 curveto +395.81367 118.1614 395.5845 117.37755 395.5845 116.27338 curveto +395.5845 112.56244 lineto +394.19388 112.56244 lineto +394.19388 110.56244 lineto +395.5845 110.56244 lineto +395.5845 108.07806 lineto +398.38138 108.07806 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +411.71732 114.914 moveto +411.71732 115.71088 lineto +405.17825 115.71088 lineto +405.24596 116.36713 405.48294 116.85932 405.88919 117.18744 curveto +406.29544 117.51557 406.86314 117.67963 407.59232 117.67963 curveto +408.18085 117.67963 408.78241 117.59369 409.397 117.42181 curveto +410.01679 117.24473 410.6522 116.97911 411.30325 116.62494 curveto +411.30325 118.78119 lineto +410.64179 119.03119 409.98033 119.21869 409.31888 119.34369 curveto +408.65741 119.4739 407.99596 119.539 407.3345 119.539 curveto +405.75117 119.539 404.5194 119.13796 403.63919 118.33588 curveto +402.76419 117.52859 402.32669 116.39838 402.32669 114.94525 curveto +402.32669 113.51817 402.75638 112.39578 403.61575 111.57806 curveto +404.48033 110.76036 405.66783 110.35151 407.17825 110.3515 curveto +408.55325 110.35151 409.6522 110.76557 410.47513 111.59369 curveto +411.30324 112.42182 411.71731 113.52859 411.71732 114.914 curveto +408.84232 113.98431 moveto +408.84231 113.45307 408.68606 113.02599 408.37357 112.70306 curveto +408.06627 112.37495 407.66262 112.21088 407.16263 112.21088 curveto +406.62096 112.21088 406.18085 112.36453 405.84232 112.67181 curveto +405.50377 112.9739 405.29283 113.4114 405.2095 113.98431 curveto +408.84232 113.98431 lineto +fill +grestore +grestore +0 0 0 setrgbcolor +[] 0 setdash +1 setlinewidth +0 setlinejoin +0 setlinecap +newpath +460.13031 263.40024 moveto +460.13031 289.97505 lineto +stroke +gsave [2.4492127e-17 0.4 -0.4 2.4492127e-17 460.13031 289.97505] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +gsave [-0.2378166 0 0 0.2269133 546.17466 132.00569] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [-0.2378166 0 0 0.2269133 546.17465 87.714359] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +2 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [-0.2378166 0 0 0.2269133 546.17465 176.29703] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [0 0.2378166 0.2269133 0 399.60191 71.056696] concat +0 0 0 setrgbcolor +[] 0 setdash +4.3013706 setlinewidth +1 setlinejoin +1 setlinecap +newpath +208.65508 282.05865 moveto +193.86388 310.15339 141.95677 326.09523 92.790977 317.64312 curveto +43.625187 309.19101 15.726964 279.5298 30.518156 251.43506 curveto +45.309349 223.34033 97.216466 207.39848 146.38226 215.85059 curveto +177.29043 221.16403 199.42278 233.82562 208.68579 251.49353 curveto +stroke +gsave [0.79891445 1.5238182 -1.5238182 0.79891445 208.68579 251.49353] concat +gsave +0 0 0 setrgbcolor +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +eofill +grestore +0 0 0 setrgbcolor +[] 0 setdash +1.25 setlinewidth +0 setlinejoin +0 setlinecap +newpath +5.77 0 moveto +-2.88 5 lineto +-2.88 -5 lineto +5.77 0 lineto +closepath +stroke +grestore +grestore +gsave [1 0 0 1 17.216929 6.5104864] concat +grestore +gsave +0 0 0 setrgbcolor +newpath +187.35507 103.05839 moveto +187.35507 104.61112 lineto +189.20566 104.61112 lineto +189.20566 105.30936 lineto +187.35507 105.30936 lineto +187.35507 108.27811 lineto +187.35507 108.72408 187.41529 109.01054 187.53574 109.13749 curveto +187.65943 109.26444 187.90846 109.32792 188.28281 109.32792 curveto +189.20566 109.32792 lineto +189.20566 110.07987 lineto +188.28281 110.07987 lineto +187.58944 110.07987 187.11093 109.95129 186.84726 109.69413 curveto +186.58359 109.43371 186.45175 108.96171 186.45175 108.27811 curveto +186.45175 105.30936 lineto +185.79257 105.30936 lineto +185.79257 104.61112 lineto +186.45175 104.61112 lineto +186.45175 103.05839 lineto +187.35507 103.05839 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +194.93808 106.77909 moveto +194.93808 110.07987 lineto +194.03964 110.07987 lineto +194.03964 106.80839 lineto +194.03964 106.29081 193.93873 105.90344 193.73691 105.64628 curveto +193.53508 105.38912 193.23235 105.26054 192.8287 105.26054 curveto +192.34368 105.26054 191.96119 105.41516 191.68124 105.7244 curveto +191.40129 106.03365 191.26132 106.4552 191.26132 106.98905 curveto +191.26132 110.07987 lineto +190.358 110.07987 lineto +190.358 102.48222 lineto +191.26132 102.48222 lineto +191.26132 105.46073 lineto +191.47616 105.13196 191.72844 104.88619 192.01816 104.72343 curveto +192.31112 104.56067 192.64804 104.47929 193.0289 104.47929 curveto +193.65715 104.47929 194.13241 104.6746 194.45468 105.06522 curveto +194.77694 105.4526 194.93807 106.02389 194.93808 106.77909 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +201.41757 107.12089 moveto +201.41757 107.56034 lineto +197.28671 107.56034 lineto +197.32577 108.17883 197.51132 108.65084 197.84335 108.97636 curveto +198.17864 109.29862 198.64413 109.45976 199.23984 109.45975 curveto +199.58489 109.45976 199.91854 109.41744 200.24081 109.3328 curveto +200.56633 109.24817 200.8886 109.12121 201.20761 108.95194 curveto +201.20761 109.80155 lineto +200.88534 109.93827 200.55494 110.04244 200.2164 110.11405 curveto +199.87785 110.18567 199.53443 110.22147 199.18613 110.22147 curveto +198.31373 110.22147 197.622 109.96757 197.11093 109.45975 curveto +196.60312 108.95194 196.34921 108.2651 196.34921 107.39921 curveto +196.34921 106.50403 196.5901 105.79439 197.07187 105.2703 curveto +197.55689 104.74296 198.20956 104.47929 199.02988 104.47929 curveto +199.76555 104.47929 200.3466 104.71692 200.77304 105.19218 curveto +201.20272 105.66419 201.41757 106.30709 201.41757 107.12089 curveto +200.51913 106.85722 moveto +200.51262 106.36568 200.37427 105.97343 200.1041 105.68046 curveto +199.83716 105.38749 199.48235 105.24101 199.03964 105.241 curveto +198.53834 105.24101 198.13632 105.38261 197.83359 105.66581 curveto +197.53411 105.94902 197.36158 106.34778 197.31601 106.8621 curveto +200.51913 106.85722 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +205.01132 105.241 moveto +204.52955 105.24101 204.14869 105.42981 203.86874 105.80741 curveto +203.58879 106.18176 203.44882 106.69609 203.44882 107.35038 curveto +203.44882 108.00468 203.58717 108.52063 203.86386 108.89823 curveto +204.14381 109.27258 204.52629 109.45976 205.01132 109.45975 curveto +205.48983 109.45976 205.86907 109.27095 206.14902 108.89335 curveto +206.42896 108.51575 206.56893 108.00142 206.56894 107.35038 curveto +206.56893 106.7026 206.42896 106.1899 206.14902 105.81229 curveto +205.86907 105.43144 205.48983 105.24101 205.01132 105.241 curveto +205.01132 104.47929 moveto +205.79257 104.47929 206.40617 104.7332 206.85214 105.241 curveto +207.2981 105.74882 207.52108 106.45195 207.52109 107.35038 curveto +207.52108 108.24556 207.2981 108.94869 206.85214 109.45975 curveto +206.40617 109.96757 205.79257 110.22147 205.01132 110.22147 curveto +204.22681 110.22147 203.61158 109.96757 203.16562 109.45975 curveto +202.72291 108.94869 202.50156 108.24556 202.50156 107.35038 curveto +202.50156 106.45195 202.72291 105.74882 203.16562 105.241 curveto +203.61158 104.7332 204.22681 104.47929 205.01132 104.47929 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +212.17441 105.45097 moveto +212.07349 105.39238 211.96282 105.35006 211.84238 105.32401 curveto +211.72519 105.29472 211.59498 105.28007 211.45175 105.28007 curveto +210.94394 105.28007 210.55331 105.44609 210.27988 105.77811 curveto +210.00969 106.10689 209.8746 106.58053 209.8746 107.19901 curveto +209.8746 110.07987 lineto +208.97128 110.07987 lineto +208.97128 104.61112 lineto +209.8746 104.61112 lineto +209.8746 105.46073 lineto +210.0634 105.12871 210.30917 104.88294 210.61191 104.72343 curveto +210.91464 104.56067 211.28248 104.47929 211.71542 104.47929 curveto +211.77727 104.47929 211.84563 104.48417 211.9205 104.49393 curveto +211.99537 104.50045 212.07838 104.51184 212.16953 104.52811 curveto +212.17441 105.45097 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +217.58945 107.12089 moveto +217.58945 107.56034 lineto +213.45859 107.56034 lineto +213.49765 108.17883 213.6832 108.65084 214.01523 108.97636 curveto +214.35051 109.29862 214.81601 109.45976 215.41171 109.45975 curveto +215.75676 109.45976 216.09042 109.41744 216.41269 109.3328 curveto +216.73821 109.24817 217.06047 109.12121 217.37949 108.95194 curveto +217.37949 109.80155 lineto +217.05722 109.93827 216.72681 110.04244 216.38828 110.11405 curveto +216.04973 110.18567 215.70631 110.22147 215.358 110.22147 curveto +214.4856 110.22147 213.79387 109.96757 213.28281 109.45975 curveto +212.77499 108.95194 212.52109 108.2651 212.52109 107.39921 curveto +212.52109 106.50403 212.76197 105.79439 213.24374 105.2703 curveto +213.72877 104.74296 214.38144 104.47929 215.20175 104.47929 curveto +215.93742 104.47929 216.51848 104.71692 216.94492 105.19218 curveto +217.3746 105.66419 217.58944 106.30709 217.58945 107.12089 curveto +216.69101 106.85722 moveto +216.68449 106.36568 216.54615 105.97343 216.27597 105.68046 curveto +216.00904 105.38749 215.65422 105.24101 215.21152 105.241 curveto +214.71021 105.24101 214.30819 105.38261 214.00546 105.66581 curveto +213.70598 105.94902 213.53346 106.34778 213.48788 106.8621 curveto +216.69101 106.85722 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +223.32187 105.66093 moveto +223.54647 105.25729 223.81503 104.95943 224.12753 104.76737 curveto +224.44003 104.57532 224.80786 104.47929 225.23105 104.47929 curveto +225.8007 104.47929 226.24016 104.67949 226.54941 105.07987 curveto +226.85864 105.47701 227.01327 106.04342 227.01328 106.77909 curveto +227.01328 110.07987 lineto +226.10995 110.07987 lineto +226.10995 106.80839 lineto +226.10995 106.2843 226.01717 105.89531 225.83163 105.6414 curveto +225.64608 105.38749 225.36288 105.26054 224.98203 105.26054 curveto +224.51652 105.26054 224.14869 105.41516 223.87851 105.7244 curveto +223.60832 106.03365 223.47323 106.4552 223.47324 106.98905 curveto +223.47324 110.07987 lineto +222.56992 110.07987 lineto +222.56992 106.80839 lineto +222.56991 106.28105 222.47714 105.89205 222.2916 105.6414 curveto +222.10604 105.38749 221.81959 105.26054 221.43222 105.26054 curveto +220.97323 105.26054 220.60865 105.41679 220.33847 105.72929 curveto +220.06829 106.03854 219.9332 106.45846 219.9332 106.98905 curveto +219.9332 110.07987 lineto +219.02988 110.07987 lineto +219.02988 104.61112 lineto +219.9332 104.61112 lineto +219.9332 105.46073 lineto +220.13827 105.12545 220.38404 104.87805 220.6705 104.71854 curveto +220.95696 104.55904 221.29713 104.47929 221.69101 104.47929 curveto +222.08814 104.47929 222.42505 104.5802 222.70175 104.78202 curveto +222.98169 104.98385 223.1884 105.27682 223.32187 105.66093 curveto +fill +grestore +gsave [1 0 0 1 17.216929 6.5104864] concat +grestore +gsave +0 0 0 setrgbcolor +newpath +470.46808 277.74594 moveto +470.46808 278.40675 470.60317 278.92596 470.87335 279.30356 curveto +471.14679 279.67791 471.52114 279.86508 471.9964 279.86508 curveto +472.47166 279.86508 472.846 279.67791 473.11945 279.30356 curveto +473.39288 278.92596 473.5296 278.40675 473.5296 277.74594 curveto +473.5296 277.08514 473.39288 276.56756 473.11945 276.19321 curveto +472.846 275.81561 472.47166 275.62681 471.9964 275.6268 curveto +471.52114 275.62681 471.14679 275.81561 470.87335 276.19321 curveto +470.60317 276.56756 470.46808 277.08514 470.46808 277.74594 curveto +473.5296 279.65512 moveto +473.3408 279.98064 473.10154 280.22315 472.81183 280.38266 curveto +472.52537 280.53891 472.18032 280.61703 471.77667 280.61703 curveto +471.11586 280.61703 470.57713 280.35336 470.16046 279.82602 curveto +469.74705 279.29868 469.54034 278.60532 469.54034 277.74594 curveto +469.54034 276.88657 469.74705 276.19321 470.16046 275.66586 curveto +470.57713 275.13852 471.11586 274.87485 471.77667 274.87485 curveto +472.18032 274.87485 472.52537 274.95461 472.81183 275.11411 curveto +473.10154 275.27036 473.3408 275.51125 473.5296 275.83676 curveto +473.5296 275.00668 lineto +474.42804 275.00668 lineto +474.42804 282.55551 lineto +473.5296 282.55551 lineto +473.5296 279.65512 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +480.95636 277.51645 moveto +480.95636 277.9559 lineto +476.8255 277.9559 lineto +476.86456 278.57439 477.05011 279.0464 477.38214 279.37192 curveto +477.71743 279.69418 478.18292 279.85532 478.77863 279.85532 curveto +479.12367 279.85532 479.45733 279.813 479.7796 279.72836 curveto +480.10512 279.64373 480.42738 279.51678 480.7464 279.3475 curveto +480.7464 280.19711 lineto +480.42413 280.33383 480.09372 280.438 479.75519 280.50961 curveto +479.41664 280.58123 479.07322 280.61703 478.72491 280.61703 curveto +477.85252 280.61703 477.16079 280.36313 476.64972 279.85532 curveto +476.14191 279.3475 475.888 278.66066 475.888 277.79477 curveto +475.888 276.89959 476.12889 276.18996 476.61066 275.66586 curveto +477.09568 275.13852 477.74835 274.87485 478.56866 274.87485 curveto +479.30434 274.87485 479.88539 275.11248 480.31183 275.58774 curveto +480.74151 276.05975 480.95635 276.70265 480.95636 277.51645 curveto +480.05792 277.25278 moveto +480.05141 276.76124 479.91306 276.36899 479.64288 276.07602 curveto +479.37595 275.78306 479.02113 275.63657 478.57843 275.63657 curveto +478.07713 275.63657 477.67511 275.77817 477.37238 276.06137 curveto +477.07289 276.34458 476.90037 276.74334 476.8548 277.25766 curveto +480.05792 277.25278 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +486.0296 275.83676 moveto +486.0296 272.87778 lineto +486.92804 272.87778 lineto +486.92804 280.47543 lineto +486.0296 280.47543 lineto +486.0296 279.65512 lineto +485.8408 279.98064 485.60154 280.22315 485.31183 280.38266 curveto +485.02537 280.53891 484.68032 280.61703 484.27667 280.61703 curveto +483.61586 280.61703 483.07713 280.35336 482.66046 279.82602 curveto +482.24705 279.29868 482.04034 278.60532 482.04034 277.74594 curveto +482.04034 276.88657 482.24705 276.19321 482.66046 275.66586 curveto +483.07713 275.13852 483.61586 274.87485 484.27667 274.87485 curveto +484.68032 274.87485 485.02537 274.95461 485.31183 275.11411 curveto +485.60154 275.27036 485.8408 275.51125 486.0296 275.83676 curveto +482.96808 277.74594 moveto +482.96808 278.40675 483.10317 278.92596 483.37335 279.30356 curveto +483.64679 279.67791 484.02114 279.86508 484.4964 279.86508 curveto +484.97166 279.86508 485.346 279.67791 485.61945 279.30356 curveto +485.89288 278.92596 486.0296 278.40675 486.0296 277.74594 curveto +486.0296 277.08514 485.89288 276.56756 485.61945 276.19321 curveto +485.346 275.81561 484.97166 275.62681 484.4964 275.6268 curveto +484.02114 275.62681 483.64679 275.81561 483.37335 276.19321 curveto +483.10317 276.56756 482.96808 277.08514 482.96808 277.74594 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +550.54895 236.85474 moveto +550.54895 237.51555 550.68404 238.03475 550.95422 238.41235 curveto +551.22766 238.7867 551.60201 238.97388 552.07727 238.97388 curveto +552.55253 238.97388 552.92688 238.7867 553.20032 238.41235 curveto +553.47375 238.03475 553.61047 237.51555 553.61047 236.85474 curveto +553.61047 236.19393 553.47375 235.67635 553.20032 235.302 curveto +552.92688 234.9244 552.55253 234.7356 552.07727 234.7356 curveto +551.60201 234.7356 551.22766 234.9244 550.95422 235.302 curveto +550.68404 235.67635 550.54895 236.19393 550.54895 236.85474 curveto +553.61047 238.76392 moveto +553.42167 239.08944 553.18241 239.33195 552.8927 239.49146 curveto +552.60624 239.64771 552.26119 239.72583 551.85754 239.72583 curveto +551.19673 239.72583 550.658 239.46216 550.24133 238.93481 curveto +549.82792 238.40747 549.62122 237.71411 549.62122 236.85474 curveto +549.62122 235.99536 549.82792 235.30201 550.24133 234.77466 curveto +550.658 234.24732 551.19673 233.98365 551.85754 233.98364 curveto +552.26119 233.98365 552.60624 234.0634 552.8927 234.2229 curveto +553.18241 234.37916 553.42167 234.62004 553.61047 234.94556 curveto +553.61047 234.11548 lineto +554.50891 234.11548 lineto +554.50891 241.66431 lineto +553.61047 241.66431 lineto +553.61047 238.76392 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +561.03723 236.62524 moveto +561.03723 237.0647 lineto +556.90637 237.0647 lineto +556.94543 237.68319 557.13098 238.15519 557.46301 238.48071 curveto +557.7983 238.80298 558.26379 238.96411 558.8595 238.96411 curveto +559.20455 238.96411 559.5382 238.92179 559.86047 238.83716 curveto +560.18599 238.75252 560.50825 238.62557 560.82727 238.4563 curveto +560.82727 239.30591 lineto +560.505 239.44263 560.1746 239.54679 559.83606 239.61841 curveto +559.49751 239.69002 559.15409 239.72583 558.80579 239.72583 curveto +557.93339 239.72583 557.24166 239.47192 556.73059 238.96411 curveto +556.22278 238.4563 555.96887 237.76945 555.96887 236.90356 curveto +555.96887 236.00839 556.20976 235.29875 556.69153 234.77466 curveto +557.17655 234.24732 557.82922 233.98365 558.64954 233.98364 curveto +559.38521 233.98365 559.96626 234.22128 560.3927 234.69653 curveto +560.82238 235.16854 561.03723 235.81145 561.03723 236.62524 curveto +560.13879 236.36157 moveto +560.13228 235.87004 559.99393 235.47779 559.72375 235.18481 curveto +559.45682 234.89185 559.10201 234.74537 558.6593 234.74536 curveto +558.158 234.74537 557.75598 234.88697 557.45325 235.17017 curveto +557.15377 235.45337 556.98124 235.85214 556.93567 236.36646 curveto +560.13879 236.36157 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.11047 234.94556 moveto +566.11047 231.98657 lineto +567.00891 231.98657 lineto +567.00891 239.58423 lineto +566.11047 239.58423 lineto +566.11047 238.76392 lineto +565.92167 239.08944 565.68241 239.33195 565.3927 239.49146 curveto +565.10624 239.64771 564.76119 239.72583 564.35754 239.72583 curveto +563.69673 239.72583 563.158 239.46216 562.74133 238.93481 curveto +562.32792 238.40747 562.12122 237.71411 562.12122 236.85474 curveto +562.12122 235.99536 562.32792 235.30201 562.74133 234.77466 curveto +563.158 234.24732 563.69673 233.98365 564.35754 233.98364 curveto +564.76119 233.98365 565.10624 234.0634 565.3927 234.2229 curveto +565.68241 234.37916 565.92167 234.62004 566.11047 234.94556 curveto +563.04895 236.85474 moveto +563.04895 237.51555 563.18404 238.03475 563.45422 238.41235 curveto +563.72766 238.7867 564.10201 238.97388 564.57727 238.97388 curveto +565.05253 238.97388 565.42688 238.7867 565.70032 238.41235 curveto +565.97375 238.03475 566.11047 237.51555 566.11047 236.85474 curveto +566.11047 236.19393 565.97375 235.67635 565.70032 235.302 curveto +565.42688 234.9244 565.05253 234.7356 564.57727 234.7356 curveto +564.10201 234.7356 563.72766 234.9244 563.45422 235.302 curveto +563.18404 235.67635 563.04895 236.19393 563.04895 236.85474 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.10266 183.66447 moveto +553.10266 184.41154 lineto +552.24329 184.41154 lineto +551.92102 184.41155 551.69641 184.47666 551.56946 184.60686 curveto +551.44576 184.73707 551.38391 184.97145 551.38391 185.30998 curveto +551.38391 185.79338 lineto +552.8634 185.79338 lineto +552.8634 186.49162 lineto +551.38391 186.49162 lineto +551.38391 191.26213 lineto +550.48059 191.26213 lineto +550.48059 186.49162 lineto +549.62122 186.49162 lineto +549.62122 185.79338 lineto +550.48059 185.79338 lineto +550.48059 185.41252 lineto +550.48059 184.8038 550.62219 184.3611 550.9054 184.0844 curveto +551.1886 183.80446 551.63782 183.66448 552.25305 183.66447 curveto +553.10266 183.66447 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.84973 185.79338 moveto +554.74817 185.79338 lineto +554.74817 191.26213 lineto +553.84973 191.26213 lineto +553.84973 185.79338 lineto +553.84973 183.66447 moveto +554.74817 183.66447 lineto +554.74817 184.80217 lineto +553.84973 184.80217 lineto +553.84973 183.66447 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +561.16907 185.79338 moveto +559.19153 188.45451 lineto +561.27161 191.26213 lineto +560.21204 191.26213 lineto +558.62024 189.11369 lineto +557.02844 191.26213 lineto +555.96887 191.26213 lineto +558.0929 188.4008 lineto +556.14954 185.79338 lineto +557.20911 185.79338 lineto +558.6593 187.74162 lineto +560.1095 185.79338 lineto +561.16907 185.79338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +552.81946 198.51311 moveto +552.09354 198.51311 551.59061 198.59612 551.31067 198.76213 curveto +551.03072 198.92815 550.89075 199.21135 550.89075 199.61174 curveto +550.89075 199.93075 550.99491 200.18466 551.20325 200.37346 curveto +551.41483 200.55901 551.70129 200.65178 552.06262 200.65178 curveto +552.56067 200.65178 552.95943 200.476 553.25891 200.12444 curveto +553.56164 199.76962 553.71301 199.29924 553.71301 198.7133 curveto +553.71301 198.51311 lineto +552.81946 198.51311 lineto +554.61145 198.14201 moveto +554.61145 201.26213 lineto +553.71301 201.26213 lineto +553.71301 200.43205 lineto +553.50793 200.76408 553.2524 201.00985 552.94641 201.16936 curveto +552.64042 201.32561 552.26607 201.40373 551.82336 201.40373 curveto +551.26347 201.40373 550.8175 201.24748 550.48547 200.93498 curveto +550.1567 200.61923 549.99231 200.19768 549.99231 199.67033 curveto +549.99231 199.0551 550.19739 198.59123 550.60754 198.27873 curveto +551.02095 197.96624 551.63619 197.80999 552.45325 197.80998 curveto +553.71301 197.80998 lineto +553.71301 197.72209 lineto +553.71301 197.30868 553.57629 196.98967 553.30286 196.76506 curveto +553.03267 196.5372 552.65181 196.42327 552.16028 196.42326 curveto +551.84778 196.42327 551.54341 196.4607 551.24719 196.53557 curveto +550.95097 196.61044 550.66614 196.72275 550.3927 196.87248 curveto +550.3927 196.0424 lineto +550.72147 195.91546 551.04049 195.82106 551.34973 195.7592 curveto +551.65897 195.6941 551.96008 195.66155 552.25305 195.66154 curveto +553.04406 195.66155 553.63488 195.86663 554.02551 196.27678 curveto +554.41613 196.68694 554.61144 197.30868 554.61145 198.14201 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +559.95325 195.95451 moveto +559.95325 196.80412 lineto +559.69934 196.67392 559.43567 196.57626 559.16223 196.51115 curveto +558.88879 196.44605 558.60559 196.4135 558.31262 196.4135 curveto +557.86666 196.4135 557.53137 196.48186 557.30676 196.61858 curveto +557.08541 196.7553 556.97473 196.96038 556.97473 197.23381 curveto +556.97473 197.44215 557.05448 197.60654 557.21399 197.72697 curveto +557.37349 197.84417 557.69413 197.95647 558.1759 198.06389 curveto +558.48352 198.13225 lineto +559.12154 198.26897 559.57401 198.46265 559.84094 198.7133 curveto +560.11112 198.9607 560.24621 199.30738 560.24622 199.75334 curveto +560.24621 200.26116 560.04439 200.66317 559.64075 200.9594 curveto +559.24035 201.25562 558.6886 201.40373 557.98547 201.40373 curveto +557.6925 201.40373 557.38651 201.37444 557.0675 201.31584 curveto +556.75175 201.2605 556.41809 201.17587 556.06653 201.06194 curveto +556.06653 200.1342 lineto +556.39856 200.30673 556.72571 200.43694 557.04797 200.52483 curveto +557.37024 200.60946 557.68925 200.65178 558.005 200.65178 curveto +558.42818 200.65178 558.7537 200.58017 558.98157 200.43694 curveto +559.20943 200.29045 559.32336 200.08537 559.32336 199.8217 curveto +559.32336 199.57756 559.24035 199.39039 559.07434 199.26018 curveto +558.91158 199.12997 558.55188 199.00465 557.99524 198.8842 curveto +557.68274 198.81096 lineto +557.1261 198.69377 556.72408 198.51474 556.47668 198.27385 curveto +556.22929 198.02971 556.10559 197.69605 556.10559 197.27287 curveto +556.10559 196.75855 556.28788 196.36142 556.65247 196.08147 curveto +557.01705 195.80152 557.53463 195.66155 558.2052 195.66154 curveto +558.53723 195.66155 558.84973 195.68596 559.1427 195.73479 curveto +559.43567 195.78362 559.70585 195.85686 559.95325 195.95451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +565.16809 195.95451 moveto +565.16809 196.80412 lineto +564.91418 196.67392 564.65051 196.57626 564.37708 196.51115 curveto +564.10363 196.44605 563.82043 196.4135 563.52747 196.4135 curveto +563.0815 196.4135 562.74621 196.48186 562.52161 196.61858 curveto +562.30025 196.7553 562.18957 196.96038 562.18958 197.23381 curveto +562.18957 197.44215 562.26933 197.60654 562.42883 197.72697 curveto +562.58834 197.84417 562.90897 197.95647 563.39075 198.06389 curveto +563.69836 198.13225 lineto +564.33638 198.26897 564.78886 198.46265 565.05579 198.7133 curveto +565.32596 198.9607 565.46105 199.30738 565.46106 199.75334 curveto +565.46105 200.26116 565.25923 200.66317 564.85559 200.9594 curveto +564.4552 201.25562 563.90344 201.40373 563.20032 201.40373 curveto +562.90735 201.40373 562.60136 201.37444 562.28235 201.31584 curveto +561.96659 201.2605 561.63293 201.17587 561.28137 201.06194 curveto +561.28137 200.1342 lineto +561.6134 200.30673 561.94055 200.43694 562.26282 200.52483 curveto +562.58508 200.60946 562.90409 200.65178 563.21985 200.65178 curveto +563.64302 200.65178 563.96854 200.58017 564.19641 200.43694 curveto +564.42427 200.29045 564.5382 200.08537 564.53821 199.8217 curveto +564.5382 199.57756 564.4552 199.39039 564.28918 199.26018 curveto +564.12642 199.12997 563.76672 199.00465 563.21008 198.8842 curveto +562.89758 198.81096 lineto +562.34094 198.69377 561.93892 198.51474 561.69153 198.27385 curveto +561.44413 198.02971 561.32043 197.69605 561.32043 197.27287 curveto +561.32043 196.75855 561.50273 196.36142 561.86731 196.08147 curveto +562.23189 195.80152 562.74947 195.66155 563.42004 195.66154 curveto +563.75207 195.66155 564.06457 195.68596 564.35754 195.73479 curveto +564.65051 195.78362 564.92069 195.85686 565.16809 195.95451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.80383 199.10393 moveto +566.80383 195.79338 lineto +567.70227 195.79338 lineto +567.70227 199.06975 lineto +567.70227 199.58733 567.80318 199.97632 568.005 200.23674 curveto +568.20683 200.4939 568.50956 200.62248 568.91321 200.62248 curveto +569.39823 200.62248 569.78072 200.46786 570.06067 200.15862 curveto +570.34387 199.84937 570.48547 199.42782 570.48547 198.89397 curveto +570.48547 195.79338 lineto +571.38391 195.79338 lineto +571.38391 201.26213 lineto +570.48547 201.26213 lineto +570.48547 200.42229 lineto +570.26737 200.75432 570.01346 201.00171 569.72375 201.16447 curveto +569.43729 201.32398 569.10363 201.40373 568.72278 201.40373 curveto +568.09452 201.40373 567.61763 201.20842 567.29211 200.81779 curveto +566.96659 200.42717 566.80383 199.85588 566.80383 199.10393 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +577.50208 196.84319 moveto +577.72668 196.43954 577.99523 196.14169 578.30774 195.94963 curveto +578.62023 195.75758 578.98807 195.66155 579.41125 195.66154 curveto +579.98091 195.66155 580.42036 195.86175 580.72961 196.26213 curveto +581.03885 196.65927 581.19347 197.22568 581.19348 197.96135 curveto +581.19348 201.26213 lineto +580.29016 201.26213 lineto +580.29016 197.99065 lineto +580.29015 197.46656 580.19738 197.07756 580.01184 196.82365 curveto +579.82629 196.56975 579.54308 196.4428 579.16223 196.44279 curveto +578.69673 196.4428 578.32889 196.59742 578.05872 196.90666 curveto +577.78853 197.21591 577.65344 197.63746 577.65344 198.17131 curveto +577.65344 201.26213 lineto +576.75012 201.26213 lineto +576.75012 197.99065 lineto +576.75012 197.46331 576.65734 197.07431 576.4718 196.82365 curveto +576.28625 196.56975 575.99979 196.4428 575.61243 196.44279 curveto +575.15344 196.4428 574.78886 196.59905 574.51868 196.91154 curveto +574.24849 197.22079 574.1134 197.64072 574.1134 198.17131 curveto +574.1134 201.26213 lineto +573.21008 201.26213 lineto +573.21008 195.79338 lineto +574.1134 195.79338 lineto +574.1134 196.64299 lineto +574.31848 196.30771 574.56425 196.06031 574.85071 195.9008 curveto +575.13716 195.7413 575.47733 195.66155 575.87122 195.66154 curveto +576.26835 195.66155 576.60526 195.76246 576.88196 195.96428 curveto +577.1619 196.16611 577.36861 196.45908 577.50208 196.84319 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +587.66809 198.30315 moveto +587.66809 198.7426 lineto +583.53723 198.7426 lineto +583.57629 199.36109 583.76184 199.8331 584.09387 200.15862 curveto +584.42916 200.48088 584.89465 200.64201 585.49036 200.64201 curveto +585.8354 200.64201 586.16906 200.5997 586.49133 200.51506 curveto +586.81685 200.43043 587.13911 200.30347 587.45813 200.1342 curveto +587.45813 200.98381 lineto +587.13586 201.12053 586.80546 201.2247 586.46692 201.29631 curveto +586.12837 201.36792 585.78495 201.40373 585.43665 201.40373 curveto +584.56425 201.40373 583.87252 201.14983 583.36145 200.64201 curveto +582.85364 200.1342 582.59973 199.44735 582.59973 198.58147 curveto +582.59973 197.68629 582.84062 196.97665 583.32239 196.45256 curveto +583.80741 195.92522 584.46008 195.66155 585.2804 195.66154 curveto +586.01607 195.66155 586.59712 195.89918 587.02356 196.37444 curveto +587.45324 196.84645 587.66809 197.48935 587.66809 198.30315 curveto +586.76965 198.03947 moveto +586.76314 197.54794 586.62479 197.15569 586.35461 196.86272 curveto +586.08768 196.56975 585.73287 196.42327 585.29016 196.42326 curveto +584.78886 196.42327 584.38684 196.56487 584.08411 196.84807 curveto +583.78463 197.13128 583.6121 197.53004 583.56653 198.04436 curveto +586.76965 198.03947 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +553.82532 147.89853 moveto +553.82532 148.60165 lineto +553.52258 148.60165 lineto +552.71203 148.60165 552.16841 148.48121 551.89172 148.24033 curveto +551.61828 147.99944 551.48156 147.5193 551.48157 146.7999 curveto +551.48157 145.6329 lineto +551.48156 145.14137 551.39367 144.8012 551.2179 144.6124 curveto +551.04211 144.4236 550.7231 144.3292 550.26086 144.32919 curveto +549.96301 144.32919 lineto +549.96301 143.63095 lineto +550.26086 143.63095 lineto +550.72636 143.63095 551.04537 143.53818 551.2179 143.35263 curveto +551.39367 143.16383 551.48156 142.82692 551.48157 142.34189 curveto +551.48157 141.17001 lineto +551.48156 140.45062 551.61828 139.9721 551.89172 139.73447 curveto +552.16841 139.49359 552.71203 139.37315 553.52258 139.37314 curveto +553.82532 139.37314 lineto +553.82532 140.07138 lineto +553.49329 140.07138 lineto +553.0343 140.07139 552.73482 140.143 552.59485 140.28622 curveto +552.45487 140.42946 552.38488 140.73057 552.38489 141.18954 curveto +552.38489 142.40048 lineto +552.38488 142.91155 552.31001 143.28265 552.16028 143.51376 curveto +552.01379 143.74489 551.76151 143.90114 551.40344 143.98251 curveto +551.76477 144.07041 552.01867 144.22991 552.16516 144.46103 curveto +552.31164 144.69215 552.38488 145.06162 552.38489 145.56943 curveto +552.38489 146.78036 lineto +552.38488 147.23935 552.45487 147.54046 552.59485 147.68369 curveto +552.73482 147.82691 553.0343 147.89853 553.49329 147.89853 curveto +553.82532 147.89853 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +559.51379 147.89853 moveto +559.85559 147.89853 lineto +560.31132 147.89853 560.60754 147.82854 560.74426 147.68857 curveto +560.88423 147.54859 560.95422 147.24586 560.95422 146.78036 curveto +560.95422 145.56943 lineto +560.95422 145.06162 561.02746 144.69215 561.17395 144.46103 curveto +561.32043 144.22991 561.57434 144.07041 561.93567 143.98251 curveto +561.57434 143.90114 561.32043 143.74489 561.17395 143.51376 curveto +561.02746 143.28265 560.95422 142.91155 560.95422 142.40048 curveto +560.95422 141.18954 lineto +560.95422 140.72731 560.88423 140.4262 560.74426 140.28622 curveto +560.60754 140.143 560.31132 140.07139 559.85559 140.07138 curveto +559.51379 140.07138 lineto +559.51379 139.37314 lineto +559.82141 139.37314 lineto +560.63196 139.37315 561.17232 139.49359 561.4425 139.73447 curveto +561.71594 139.9721 561.85266 140.45062 561.85266 141.17001 curveto +561.85266 142.34189 lineto +561.85266 142.82692 561.94055 143.16383 562.11633 143.35263 curveto +562.29211 143.53818 562.61112 143.63095 563.07336 143.63095 curveto +563.3761 143.63095 lineto +563.3761 144.32919 lineto +563.07336 144.32919 lineto +562.61112 144.3292 562.29211 144.4236 562.11633 144.6124 curveto +561.94055 144.8012 561.85266 145.14137 561.85266 145.6329 curveto +561.85266 146.7999 lineto +561.85266 147.5193 561.71594 147.99944 561.4425 148.24033 curveto +561.17232 148.48121 560.63196 148.60165 559.82141 148.60165 curveto +559.51379 148.60165 lineto +559.51379 147.89853 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +554.20129 153.67001 moveto +554.20129 156.97079 lineto +553.30286 156.97079 lineto +553.30286 153.69931 lineto +553.30285 153.18174 553.20194 152.79437 553.00012 152.5372 curveto +552.7983 152.28004 552.49556 152.15146 552.09192 152.15146 curveto +551.60689 152.15146 551.2244 152.30609 550.94446 152.61533 curveto +550.66451 152.92457 550.52453 153.34612 550.52454 153.87997 curveto +550.52454 156.97079 lineto +549.62122 156.97079 lineto +549.62122 151.50204 lineto +550.52454 151.50204 lineto +550.52454 152.35165 lineto +550.73938 152.02288 550.99166 151.77711 551.28137 151.61435 curveto +551.57434 151.45159 551.91125 151.37021 552.29211 151.37021 curveto +552.92037 151.37021 553.39563 151.56553 553.7179 151.95615 curveto +554.04016 152.34352 554.20129 152.91481 554.20129 153.67001 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +560.68079 154.01181 moveto +560.68079 154.45126 lineto +556.54993 154.45126 lineto +556.58899 155.06975 556.77453 155.54176 557.10657 155.86728 curveto +557.44185 156.18955 557.90735 156.35068 558.50305 156.35068 curveto +558.8481 156.35068 559.18176 156.30836 559.50403 156.22372 curveto +559.82954 156.13909 560.15181 156.01214 560.47083 155.84286 curveto +560.47083 156.69247 lineto +560.14855 156.82919 559.81815 156.93336 559.47961 157.00497 curveto +559.14107 157.07659 558.79764 157.1124 558.44934 157.1124 curveto +557.57694 157.1124 556.88521 156.85849 556.37415 156.35068 curveto +555.86633 155.84287 555.61243 155.15602 555.61243 154.29013 curveto +555.61243 153.39495 555.85331 152.68532 556.33508 152.16122 curveto +556.82011 151.63389 557.47278 151.37021 558.29309 151.37021 curveto +559.02876 151.37021 559.60982 151.60784 560.03625 152.0831 curveto +560.46594 152.55511 560.68078 153.19801 560.68079 154.01181 curveto +559.78235 153.74814 moveto +559.77583 153.25661 559.63749 152.86435 559.36731 152.57138 curveto +559.10038 152.27842 558.74556 152.13193 558.30286 152.13193 curveto +557.80155 152.13193 557.39953 152.27353 557.0968 152.55673 curveto +556.79732 152.83994 556.62479 153.2387 556.57922 153.75302 curveto +559.78235 153.74814 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +566.52551 151.50204 moveto +564.54797 154.16318 lineto +566.62805 156.97079 lineto +565.56848 156.97079 lineto +563.97668 154.82236 lineto +562.38489 156.97079 lineto +561.32532 156.97079 lineto +563.44934 154.10947 lineto +561.50598 151.50204 lineto +562.56555 151.50204 lineto +564.01575 153.45029 lineto +565.46594 151.50204 lineto +566.52551 151.50204 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +568.78625 149.94931 moveto +568.78625 151.50204 lineto +570.63684 151.50204 lineto +570.63684 152.20029 lineto +568.78625 152.20029 lineto +568.78625 155.16904 lineto +568.78625 155.615 568.84647 155.90146 568.96692 156.02841 curveto +569.09061 156.15537 569.33964 156.21884 569.71399 156.21884 curveto +570.63684 156.21884 lineto +570.63684 156.97079 lineto +569.71399 156.97079 lineto +569.02063 156.97079 568.54211 156.84221 568.27844 156.58505 curveto +568.01477 156.32464 567.88293 155.85263 567.88293 155.16904 curveto +567.88293 152.20029 lineto +567.22375 152.20029 lineto +567.22375 151.50204 lineto +567.88293 151.50204 lineto +567.88293 149.94931 lineto +568.78625 149.94931 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +483.33514 94.963516 moveto +483.33514 98.264297 lineto +482.43671 98.264297 lineto +482.43671 94.992813 lineto +482.4367 94.475239 482.33579 94.087869 482.13397 93.830704 curveto +481.93215 93.573547 481.62941 93.444966 481.22577 93.444962 curveto +480.74074 93.444966 480.35825 93.599589 480.07831 93.908829 curveto +479.79836 94.218078 479.65838 94.639627 479.65839 95.173477 curveto +479.65839 98.264297 lineto +478.75507 98.264297 lineto +478.75507 92.795547 lineto +479.65839 92.795547 lineto +479.65839 93.645157 lineto +479.87323 93.316386 480.12551 93.070618 480.41522 92.907852 curveto +480.70819 92.745097 481.0451 92.663717 481.42596 92.663712 curveto +482.05422 92.663717 482.52948 92.859029 482.85175 93.249649 curveto +483.17401 93.637023 483.33514 94.208312 483.33514 94.963516 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +487.25604 93.42543 moveto +486.77427 93.425435 486.39341 93.614237 486.11346 93.991837 curveto +485.83351 94.366189 485.69354 94.880512 485.69354 95.534805 curveto +485.69354 96.189104 485.83189 96.705054 486.10858 97.082657 curveto +486.38853 97.457007 486.77101 97.644181 487.25604 97.64418 curveto +487.73455 97.644181 488.11379 97.455379 488.39374 97.077774 curveto +488.67368 96.700171 488.81366 96.185849 488.81366 95.534805 curveto +488.81366 94.887022 488.67368 94.374327 488.39374 93.996719 curveto +488.11379 93.615865 487.73455 93.425435 487.25604 93.42543 curveto +487.25604 92.663712 moveto +488.03729 92.663717 488.65089 92.917623 489.09686 93.42543 curveto +489.54282 93.933247 489.7658 94.636371 489.76581 95.534805 curveto +489.7658 96.429989 489.54282 97.133114 489.09686 97.64418 curveto +488.65089 98.151993 488.03729 98.405899 487.25604 98.405899 curveto +486.47153 98.405899 485.8563 98.151993 485.41034 97.64418 curveto +484.96763 97.133114 484.74628 96.429989 484.74628 95.534805 curveto +484.74628 94.636371 484.96763 93.933247 485.41034 93.42543 curveto +485.8563 92.917623 486.47153 92.663717 487.25604 92.663712 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +492.13885 91.242813 moveto +492.13885 92.795547 lineto +493.98944 92.795547 lineto +493.98944 93.49379 lineto +492.13885 93.49379 lineto +492.13885 96.46254 lineto +492.13885 96.908505 492.19907 97.194963 492.31952 97.321915 curveto +492.44321 97.448869 492.69224 97.512345 493.06659 97.512344 curveto +493.98944 97.512344 lineto +493.98944 98.264297 lineto +493.06659 98.264297 lineto +492.37323 98.264297 491.89471 98.135717 491.63104 97.878555 curveto +491.36737 97.618139 491.23553 97.146135 491.23553 96.46254 curveto +491.23553 93.49379 lineto +490.57635 93.49379 lineto +490.57635 92.795547 lineto +491.23553 92.795547 lineto +491.23553 91.242813 lineto +492.13885 91.242813 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +499.8537 95.305313 moveto +499.8537 95.744766 lineto +495.72284 95.744766 lineto +495.7619 96.363258 495.94745 96.835262 496.27948 97.160782 curveto +496.61476 97.483048 497.08026 97.644181 497.67596 97.64418 curveto +498.02101 97.644181 498.35467 97.601863 498.67694 97.517227 curveto +499.00246 97.432593 499.32472 97.30564 499.64374 97.136368 curveto +499.64374 97.985977 lineto +499.32147 98.122696 498.99106 98.226863 498.65253 98.298477 curveto +498.31398 98.370092 497.97056 98.405899 497.62225 98.405899 curveto +496.74986 98.405899 496.05812 98.151993 495.54706 97.64418 curveto +495.03924 97.136369 494.78534 96.449521 494.78534 95.583633 curveto +494.78534 94.688455 495.02622 93.97882 495.508 93.454727 curveto +495.99302 92.927389 496.64569 92.663717 497.466 92.663712 curveto +498.20168 92.663717 498.78273 92.901347 499.20917 93.376602 curveto +499.63885 93.848612 499.85369 94.491515 499.8537 95.305313 curveto +498.95526 95.041641 moveto +498.94875 94.550108 498.8104 94.157856 498.54022 93.864883 curveto +498.27329 93.571919 497.91847 93.425435 497.47577 93.42543 curveto +496.97446 93.425435 496.57245 93.567037 496.26971 93.850235 curveto +495.97023 94.133442 495.79771 94.532205 495.75214 95.046524 curveto +498.95526 95.041641 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +478.78925 100.66664 moveto +479.68768 100.66664 lineto +479.68768 108.2643 lineto +478.78925 108.2643 lineto +478.78925 100.66664 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +486.24042 105.30531 moveto +486.24042 105.74477 lineto +482.10956 105.74477 lineto +482.14862 106.36326 482.33417 106.83526 482.6662 107.16078 curveto +483.00148 107.48305 483.46698 107.64418 484.06268 107.64418 curveto +484.40773 107.64418 484.74139 107.60186 485.06366 107.51723 curveto +485.38918 107.43259 485.71144 107.30564 486.03046 107.13637 curveto +486.03046 107.98598 lineto +485.70819 108.1227 485.37778 108.22686 485.03925 108.29848 curveto +484.7007 108.37009 484.35728 108.4059 484.00897 108.4059 curveto +483.13657 108.4059 482.44484 108.15199 481.93378 107.64418 curveto +481.42596 107.13637 481.17206 106.44952 481.17206 105.58363 curveto +481.17206 104.68845 481.41294 103.97882 481.89471 103.45473 curveto +482.37974 102.92739 483.03241 102.66372 483.85272 102.66371 curveto +484.5884 102.66372 485.16945 102.90135 485.59589 103.3766 curveto +486.02557 103.84861 486.24041 104.49151 486.24042 105.30531 curveto +485.34198 105.04164 moveto +485.33546 104.55011 485.19712 104.15786 484.92694 103.86488 curveto +484.66001 103.57192 484.30519 103.42544 483.86249 103.42543 curveto +483.36118 103.42544 482.95917 103.56704 482.65643 103.85023 curveto +482.35695 104.13344 482.18443 104.5322 482.13885 105.04652 curveto +485.34198 105.04164 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +488.6037 101.24281 moveto +488.6037 102.79555 lineto +490.45428 102.79555 lineto +490.45428 103.49379 lineto +488.6037 103.49379 lineto +488.6037 106.46254 lineto +488.6037 106.9085 488.66392 107.19496 488.78436 107.32191 curveto +488.90806 107.44887 489.15708 107.51235 489.53143 107.51234 curveto +490.45428 107.51234 lineto +490.45428 108.2643 lineto +489.53143 108.2643 lineto +488.83807 108.2643 488.35956 108.13572 488.09589 107.87856 curveto +487.83221 107.61814 487.70038 107.14613 487.70038 106.46254 curveto +487.70038 103.49379 lineto +487.0412 103.49379 lineto +487.0412 102.79555 lineto +487.70038 102.79555 lineto +487.70038 101.24281 lineto +488.6037 101.24281 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +44.641342 188.13469 moveto +44.641342 184.82414 lineto +45.53978 184.82414 lineto +45.53978 188.10051 lineto +45.539778 188.61809 45.640689 189.00709 45.842514 189.2675 curveto +46.044335 189.52466 46.347069 189.65324 46.750717 189.65324 curveto +47.23574 189.65324 47.618226 189.49862 47.898178 189.18938 curveto +48.181377 188.88013 48.322978 188.45858 48.322983 187.92473 curveto +48.322983 184.82414 lineto +49.22142 184.82414 lineto +49.22142 190.29289 lineto +48.322983 190.29289 lineto +48.322983 189.45305 lineto +48.10488 189.78508 47.850974 190.03248 47.561264 190.19524 curveto +47.274802 190.35474 46.941144 190.43449 46.560287 190.43449 curveto +45.93203 190.43449 45.455143 190.23918 45.129623 189.84856 curveto +44.804102 189.45793 44.641341 188.88664 44.641342 188.13469 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +54.5681 184.98528 moveto +54.5681 185.83488 lineto +54.31419 185.70468 54.050518 185.60702 53.777084 185.54192 curveto +53.503643 185.47682 53.220441 185.44426 52.927475 185.44426 curveto +52.481509 185.44426 52.146223 185.51262 51.921616 185.64934 curveto +51.70026 185.78606 51.589583 185.99114 51.589584 186.26457 curveto +51.589583 186.47291 51.669335 186.6373 51.828842 186.75774 curveto +51.988346 186.87493 52.308983 186.98723 52.790756 187.09465 curveto +53.098373 187.16301 lineto +53.736391 187.29973 54.188864 187.49342 54.455795 187.74406 curveto +54.725973 187.99146 54.861064 188.33814 54.861069 188.7841 curveto +54.861064 189.29192 54.659241 189.69393 54.2556 189.99016 curveto +53.855206 190.28638 53.303448 190.43449 52.600327 190.43449 curveto +52.307356 190.43449 52.001366 190.4052 51.682358 190.3466 curveto +51.366601 190.29126 51.032943 190.20663 50.681381 190.0927 curveto +50.681381 189.16496 lineto +51.013412 189.33749 51.34056 189.4677 51.662827 189.55559 curveto +51.98509 189.64022 52.3041 189.68254 52.619858 189.68254 curveto +53.043032 189.68254 53.368552 189.61093 53.59642 189.4677 curveto +53.824281 189.32121 53.938213 189.11614 53.938217 188.85246 curveto +53.938213 188.60832 53.855206 188.42115 53.689194 188.29094 curveto +53.52643 188.16073 53.16673 188.03541 52.610092 187.91496 curveto +52.297592 187.84172 lineto +51.74095 187.72454 51.338932 187.5455 51.091537 187.30461 curveto +50.844141 187.06047 50.720443 186.72682 50.720444 186.30363 curveto +50.720443 185.78932 50.902735 185.39218 51.267319 185.11223 curveto +51.631901 184.83229 52.149478 184.69231 52.820053 184.69231 curveto +53.152081 184.69231 53.464581 184.71673 53.757553 184.76555 curveto +54.050518 184.81438 54.3207 184.88762 54.5681 184.98528 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +56.296616 184.82414 moveto +57.195053 184.82414 lineto +57.195053 190.29289 lineto +56.296616 190.29289 lineto +56.296616 184.82414 lineto +56.296616 182.69524 moveto +57.195053 182.69524 lineto +57.195053 183.83293 lineto +56.296616 183.83293 lineto +56.296616 182.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +63.615952 186.99211 moveto +63.615952 190.29289 lineto +62.717514 190.29289 lineto +62.717514 187.02141 lineto +62.717509 186.50383 62.616598 186.11646 62.41478 185.8593 curveto +62.212953 185.60214 61.910219 185.47356 61.506577 185.47356 curveto +61.021548 185.47356 60.639061 185.62818 60.359116 185.93742 curveto +60.079166 186.24667 59.939192 186.66822 59.939194 187.20207 curveto +59.939194 190.29289 lineto +59.035873 190.29289 lineto +59.035873 184.82414 lineto +59.939194 184.82414 lineto +59.939194 185.67375 lineto +60.154035 185.34498 60.406314 185.09921 60.69603 184.93645 curveto +60.988996 184.77369 61.325909 184.69231 61.706772 184.69231 curveto +62.335023 184.69231 62.810283 184.88762 63.132553 185.27824 curveto +63.454813 185.66562 63.615946 186.23691 63.615952 186.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +69.016342 187.49504 moveto +69.016338 186.844 68.881247 186.33945 68.611069 185.98137 curveto +68.344138 185.6233 67.968162 185.44426 67.483139 185.44426 curveto +67.001366 185.44426 66.625389 185.6233 66.355209 185.98137 curveto +66.088281 186.33945 65.954817 186.844 65.954819 187.49504 curveto +65.954817 188.14283 66.088281 188.64576 66.355209 189.00383 curveto +66.625389 189.3619 67.001366 189.54094 67.483139 189.54094 curveto +67.968162 189.54094 68.344138 189.3619 68.611069 189.00383 curveto +68.881247 188.64576 69.016338 188.14283 69.016342 187.49504 curveto +69.91478 189.61418 moveto +69.914774 190.54517 69.708069 191.2369 69.294662 191.68938 curveto +68.881247 192.1451 68.248109 192.37297 67.395248 192.37297 curveto +67.079491 192.37297 66.781639 192.34855 66.501694 192.29973 curveto +66.221744 192.25415 65.949934 192.18254 65.686264 192.08488 curveto +65.686264 191.21086 lineto +65.949934 191.35409 66.210351 191.45988 66.467514 191.52824 curveto +66.724673 191.5966 66.986717 191.63078 67.253647 191.63078 curveto +67.842836 191.63078 68.283916 191.47616 68.576889 191.16692 curveto +68.869853 190.86093 69.016338 190.39706 69.016342 189.77531 curveto +69.016342 189.33098 lineto +68.830791 189.65324 68.593161 189.89413 68.303452 190.05363 curveto +68.013734 190.21314 67.667055 190.29289 67.263412 190.29289 curveto +66.592837 190.29289 66.052473 190.03736 65.642319 189.52629 curveto +65.232162 189.01522 65.027084 188.33814 65.027084 187.49504 curveto +65.027084 186.64869 65.232162 185.96998 65.642319 185.45891 curveto +66.052473 184.94785 66.592837 184.69231 67.263412 184.69231 curveto +67.667055 184.69231 68.013734 184.77206 68.303452 184.93156 curveto +68.593161 185.09107 68.830791 185.33196 69.016342 185.65422 curveto +69.016342 184.82414 lineto +69.91478 184.82414 lineto +69.91478 189.61418 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +44.641342 198.13469 moveto +44.641342 194.82414 lineto +45.53978 194.82414 lineto +45.53978 198.10051 lineto +45.539778 198.61809 45.640689 199.00709 45.842514 199.2675 curveto +46.044335 199.52466 46.347069 199.65324 46.750717 199.65324 curveto +47.23574 199.65324 47.618226 199.49862 47.898178 199.18938 curveto +48.181377 198.88013 48.322978 198.45858 48.322983 197.92473 curveto +48.322983 194.82414 lineto +49.22142 194.82414 lineto +49.22142 200.29289 lineto +48.322983 200.29289 lineto +48.322983 199.45305 lineto +48.10488 199.78508 47.850974 200.03248 47.561264 200.19524 curveto +47.274802 200.35474 46.941144 200.43449 46.560287 200.43449 curveto +45.93203 200.43449 45.455143 200.23918 45.129623 199.84856 curveto +44.804102 199.45793 44.641341 198.88664 44.641342 198.13469 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +55.62767 196.99211 moveto +55.62767 200.29289 lineto +54.729233 200.29289 lineto +54.729233 197.02141 lineto +54.729228 196.50383 54.628317 196.11646 54.426498 195.8593 curveto +54.224671 195.60214 53.921937 195.47356 53.518295 195.47356 curveto +53.033266 195.47356 52.65078 195.62818 52.370834 195.93742 curveto +52.090884 196.24667 51.950911 196.66822 51.950912 197.20207 curveto +51.950912 200.29289 lineto +51.047592 200.29289 lineto +51.047592 194.82414 lineto +51.950912 194.82414 lineto +51.950912 195.67375 lineto +52.165754 195.34498 52.418033 195.09921 52.707748 194.93645 curveto +53.000714 194.77369 53.337628 194.69231 53.718491 194.69231 curveto +54.346742 194.69231 54.822002 194.88762 55.144272 195.27824 curveto +55.466532 195.66562 55.627665 196.23691 55.62767 196.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +60.197983 192.69524 moveto +60.197983 193.44231 lineto +59.338608 193.44231 lineto +59.01634 193.44231 58.79173 193.50742 58.66478 193.63762 curveto +58.54108 193.76783 58.479231 194.00221 58.479233 194.34074 curveto +58.479233 194.82414 lineto +59.958725 194.82414 lineto +59.958725 195.52238 lineto +58.479233 195.52238 lineto +58.479233 200.29289 lineto +57.575912 200.29289 lineto +57.575912 195.52238 lineto +56.716537 195.52238 lineto +56.716537 194.82414 lineto +57.575912 194.82414 lineto +57.575912 194.44328 lineto +57.575911 193.83457 57.717513 193.39186 58.000717 193.11516 curveto +58.283918 192.83522 58.733137 192.69524 59.348373 192.69524 curveto +60.197983 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +63.064194 195.45403 moveto +62.58242 195.45403 62.201561 195.64283 61.921616 196.02043 curveto +61.641666 196.39478 61.501692 196.90911 61.501694 197.5634 curveto +61.501692 198.2177 61.640038 198.73365 61.916733 199.11125 curveto +62.196679 199.4856 62.579165 199.67278 63.064194 199.67278 curveto +63.542706 199.67278 63.921937 199.48397 64.201889 199.10637 curveto +64.481832 198.72877 64.621806 198.21444 64.621811 197.5634 curveto +64.621806 196.91562 64.481832 196.40292 64.201889 196.02531 curveto +63.921937 195.64446 63.542706 195.45403 63.064194 195.45403 curveto +63.064194 194.69231 moveto +63.84544 194.69231 64.459046 194.94622 64.905014 195.45403 curveto +65.350972 195.96184 65.573954 196.66497 65.573959 197.5634 curveto +65.573954 198.45858 65.350972 199.16171 64.905014 199.67278 curveto +64.459046 200.18059 63.84544 200.43449 63.064194 200.43449 curveto +62.279686 200.43449 61.664452 200.18059 61.218491 199.67278 curveto +60.775781 199.16171 60.554428 198.45858 60.554428 197.5634 curveto +60.554428 196.66497 60.775781 195.96184 61.218491 195.45403 curveto +61.664452 194.94622 62.279686 194.69231 63.064194 194.69231 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +67.058334 192.69524 moveto +67.956772 192.69524 lineto +67.956772 200.29289 lineto +67.058334 200.29289 lineto +67.058334 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +73.430405 195.65422 moveto +73.430405 192.69524 lineto +74.328842 192.69524 lineto +74.328842 200.29289 lineto +73.430405 200.29289 lineto +73.430405 199.47258 lineto +73.241598 199.7981 73.002341 200.04061 72.712631 200.20012 curveto +72.426169 200.35637 72.081118 200.43449 71.677475 200.43449 curveto +71.016666 200.43449 70.477929 200.17082 70.061264 199.64348 curveto +69.647852 199.11614 69.441146 198.42278 69.441147 197.5634 curveto +69.441146 196.70403 69.647852 196.01067 70.061264 195.48332 curveto +70.477929 194.95598 71.016666 194.69231 71.677475 194.69231 curveto +72.081118 194.69231 72.426169 194.77206 72.712631 194.93156 curveto +73.002341 195.08782 73.241598 195.3287 73.430405 195.65422 curveto +70.368881 197.5634 moveto +70.36888 198.22421 70.503971 198.74341 70.774155 199.12102 curveto +71.04759 199.49537 71.421939 199.68254 71.897202 199.68254 curveto +72.372458 199.68254 72.746807 199.49537 73.020248 199.12102 curveto +73.293682 198.74341 73.4304 198.22421 73.430405 197.5634 curveto +73.4304 196.9026 73.293682 196.38502 73.020248 196.01067 curveto +72.746807 195.63307 72.372458 195.44426 71.897202 195.44426 curveto +71.421939 195.44426 71.04759 195.63307 70.774155 196.01067 curveto +70.503971 196.38502 70.36888 196.9026 70.368881 197.5634 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +76.179428 194.82414 moveto +77.077866 194.82414 lineto +77.077866 200.29289 lineto +76.179428 200.29289 lineto +76.179428 194.82414 lineto +76.179428 192.69524 moveto +77.077866 192.69524 lineto +77.077866 193.83293 lineto +76.179428 193.83293 lineto +76.179428 192.69524 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +83.498764 196.99211 moveto +83.498764 200.29289 lineto +82.600327 200.29289 lineto +82.600327 197.02141 lineto +82.600322 196.50383 82.499411 196.11646 82.297592 195.8593 curveto +82.095765 195.60214 81.793031 195.47356 81.389389 195.47356 curveto +80.90436 195.47356 80.521874 195.62818 80.241928 195.93742 curveto +79.961978 196.24667 79.822004 196.66822 79.822006 197.20207 curveto +79.822006 200.29289 lineto +78.918686 200.29289 lineto +78.918686 194.82414 lineto +79.822006 194.82414 lineto +79.822006 195.67375 lineto +80.036848 195.34498 80.289126 195.09921 80.578842 194.93645 curveto +80.871808 194.77369 81.208722 194.69231 81.589584 194.69231 curveto +82.217835 194.69231 82.693095 194.88762 83.015366 195.27824 curveto +83.337626 195.66562 83.498759 196.23691 83.498764 196.99211 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +88.899155 197.49504 moveto +88.89915 196.844 88.764059 196.33945 88.493881 195.98137 curveto +88.22695 195.6233 87.850974 195.44426 87.365952 195.44426 curveto +86.884178 195.44426 86.508202 195.6233 86.238022 195.98137 curveto +85.971093 196.33945 85.83763 196.844 85.837631 197.49504 curveto +85.83763 198.14283 85.971093 198.64576 86.238022 199.00383 curveto +86.508202 199.3619 86.884178 199.54094 87.365952 199.54094 curveto +87.850974 199.54094 88.22695 199.3619 88.493881 199.00383 curveto +88.764059 198.64576 88.89915 198.14283 88.899155 197.49504 curveto +89.797592 199.61418 moveto +89.797587 200.54517 89.590881 201.2369 89.177475 201.68938 curveto +88.764059 202.1451 88.130922 202.37297 87.278061 202.37297 curveto +86.962303 202.37297 86.664452 202.34855 86.384506 202.29973 curveto +86.104557 202.25415 85.832747 202.18254 85.569077 202.08488 curveto +85.569077 201.21086 lineto +85.832747 201.35409 86.093163 201.45988 86.350327 201.52824 curveto +86.607486 201.5966 86.86953 201.63078 87.136459 201.63078 curveto +87.725649 201.63078 88.166729 201.47616 88.459702 201.16692 curveto +88.752666 200.86093 88.89915 200.39706 88.899155 199.77531 curveto +88.899155 199.33098 lineto +88.713603 199.65324 88.475973 199.89413 88.186264 200.05363 curveto +87.896547 200.21314 87.549868 200.29289 87.146225 200.29289 curveto +86.47565 200.29289 85.935286 200.03736 85.525131 199.52629 curveto +85.114974 199.01522 84.909896 198.33814 84.909897 197.49504 curveto +84.909896 196.64869 85.114974 195.96998 85.525131 195.45891 curveto +85.935286 194.94785 86.47565 194.69231 87.146225 194.69231 curveto +87.549868 194.69231 87.896547 194.77206 88.186264 194.93156 curveto +88.475973 195.09107 88.713603 195.33196 88.899155 195.65422 curveto +88.899155 194.82414 lineto +89.797592 194.82414 lineto +89.797592 199.61418 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +380.48996 223.50369 moveto +380.48996 225.05643 lineto +382.34055 225.05643 lineto +382.34055 225.75467 lineto +380.48996 225.75467 lineto +380.48996 228.72342 lineto +380.48996 229.16938 380.55018 229.45584 380.67062 229.58279 curveto +380.79432 229.70975 381.04334 229.77322 381.41769 229.77322 curveto +382.34055 229.77322 lineto +382.34055 230.52518 lineto +381.41769 230.52518 lineto +380.72433 230.52518 380.24582 230.3966 379.98215 230.13943 curveto +379.71847 229.87902 379.58664 229.40701 379.58664 228.72342 curveto +379.58664 225.75467 lineto +378.92746 225.75467 lineto +378.92746 225.05643 lineto +379.58664 225.05643 lineto +379.58664 223.50369 lineto +380.48996 223.50369 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +388.07297 227.2244 moveto +388.07297 230.52518 lineto +387.17453 230.52518 lineto +387.17453 227.25369 lineto +387.17453 226.73612 387.07361 226.34875 386.8718 226.09158 curveto +386.66997 225.83443 386.36723 225.70585 385.96359 225.70584 curveto +385.47856 225.70585 385.09608 225.86047 384.81613 226.16971 curveto +384.53618 226.47896 384.39621 226.90051 384.39621 227.43436 curveto +384.39621 230.52518 lineto +383.49289 230.52518 lineto +383.49289 222.92752 lineto +384.39621 222.92752 lineto +384.39621 225.90604 lineto +384.61105 225.57727 384.86333 225.3315 385.15305 225.16873 curveto +385.44601 225.00598 385.78293 224.9246 386.16379 224.92459 curveto +386.79204 224.9246 387.2673 225.11991 387.58957 225.51053 curveto +387.91183 225.8979 388.07296 226.46919 388.07297 227.2244 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +394.55246 227.56619 moveto +394.55246 228.00565 lineto +390.4216 228.00565 lineto +390.46066 228.62414 390.64621 229.09614 390.97824 229.42166 curveto +391.31353 229.74393 391.77902 229.90506 392.37473 229.90506 curveto +392.71977 229.90506 393.05343 229.86274 393.3757 229.77811 curveto +393.70122 229.69347 394.02348 229.56652 394.3425 229.39725 curveto +394.3425 230.24686 lineto +394.02023 230.38358 393.68982 230.48774 393.35129 230.55936 curveto +393.01274 230.63097 392.66932 230.66678 392.32101 230.66678 curveto +391.44862 230.66678 390.75688 230.41287 390.24582 229.90506 curveto +389.73801 229.39725 389.4841 228.7104 389.4841 227.84451 curveto +389.4841 226.94933 389.72498 226.2397 390.20676 225.71561 curveto +390.69178 225.18827 391.34445 224.9246 392.16476 224.92459 curveto +392.90044 224.9246 393.48149 225.16223 393.90793 225.63748 curveto +394.33761 226.10949 394.55245 226.75239 394.55246 227.56619 curveto +393.65402 227.30252 moveto +393.64751 226.81099 393.50916 226.41874 393.23898 226.12576 curveto +392.97205 225.8328 392.61723 225.68631 392.17453 225.68631 curveto +391.67323 225.68631 391.27121 225.82792 390.96848 226.11111 curveto +390.66899 226.39432 390.49647 226.79308 390.4509 227.3074 curveto +393.65402 227.30252 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +400.57297 227.2244 moveto +400.57297 230.52518 lineto +399.67453 230.52518 lineto +399.67453 227.25369 lineto +399.67453 226.73612 399.57361 226.34875 399.3718 226.09158 curveto +399.16997 225.83443 398.86723 225.70585 398.46359 225.70584 curveto +397.97856 225.70585 397.59608 225.86047 397.31613 226.16971 curveto +397.03618 226.47896 396.89621 226.90051 396.89621 227.43436 curveto +396.89621 230.52518 lineto +395.99289 230.52518 lineto +395.99289 225.05643 lineto +396.89621 225.05643 lineto +396.89621 225.90604 lineto +397.11105 225.57727 397.36333 225.3315 397.65305 225.16873 curveto +397.94601 225.00598 398.28293 224.9246 398.66379 224.92459 curveto +399.29204 224.9246 399.7673 225.11991 400.08957 225.51053 curveto +400.41183 225.8979 400.57296 226.46919 400.57297 227.2244 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +239.47623 229.75269 moveto +239.47623 233.05347 lineto +238.57779 233.05347 lineto +238.57779 229.78198 lineto +238.57778 229.26441 238.47687 228.87704 238.27505 228.61987 curveto +238.07323 228.36272 237.77049 228.23414 237.36685 228.23413 curveto +236.88182 228.23414 236.49934 228.38876 236.21939 228.698 curveto +235.93944 229.00725 235.79947 229.4288 235.79947 229.96265 curveto +235.79947 233.05347 lineto +234.89615 233.05347 lineto +234.89615 225.45581 lineto +235.79947 225.45581 lineto +235.79947 228.43433 lineto +236.01431 228.10556 236.26659 227.85979 236.5563 227.69702 curveto +236.84927 227.53427 237.18618 227.45289 237.56705 227.45288 curveto +238.1953 227.45289 238.67056 227.6482 238.99283 228.03882 curveto +239.31509 228.42619 239.47622 228.99748 239.47623 229.75269 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +243.76334 230.30444 moveto +243.03742 230.30445 242.53449 230.38745 242.25455 230.55347 curveto +241.9746 230.71948 241.83462 231.00269 241.83463 231.40308 curveto +241.83462 231.72209 241.93879 231.97599 242.14713 232.16479 curveto +242.35871 232.35034 242.64517 232.44312 243.0065 232.44312 curveto +243.50454 232.44312 243.90331 232.26733 244.20279 231.91577 curveto +244.50552 231.56096 244.65689 231.09058 244.65689 230.50464 curveto +244.65689 230.30444 lineto +243.76334 230.30444 lineto +245.55533 229.93335 moveto +245.55533 233.05347 lineto +244.65689 233.05347 lineto +244.65689 232.22339 lineto +244.45181 232.55542 244.19628 232.80119 243.89029 232.96069 curveto +243.5843 233.11694 243.20995 233.19507 242.76724 233.19507 curveto +242.20734 233.19507 241.76138 233.03882 241.42935 232.72632 curveto +241.10057 232.41056 240.93619 231.98901 240.93619 231.46167 curveto +240.93619 230.84644 241.14127 230.38257 241.55142 230.07007 curveto +241.96483 229.75757 242.58007 229.60132 243.39713 229.60132 curveto +244.65689 229.60132 lineto +244.65689 229.51343 lineto +244.65689 229.10002 244.52017 228.78101 244.24673 228.5564 curveto +243.97655 228.32854 243.59569 228.2146 243.10416 228.2146 curveto +242.79165 228.2146 242.48729 228.25204 242.19107 228.3269 curveto +241.89485 228.40178 241.61001 228.51408 241.33658 228.66382 curveto +241.33658 227.83374 lineto +241.66535 227.70679 241.98436 227.61239 242.29361 227.55054 curveto +242.60285 227.48544 242.90396 227.45289 243.19693 227.45288 curveto +243.98794 227.45289 244.57876 227.65796 244.96939 228.06812 curveto +245.36001 228.47828 245.55532 229.10002 245.55533 229.93335 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +246.76627 227.58472 moveto +247.71841 227.58472 lineto +249.4274 232.17456 lineto +251.13638 227.58472 lineto +252.08853 227.58472 lineto +250.03775 233.05347 lineto +248.81705 233.05347 lineto +246.76627 227.58472 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +258.0065 230.09448 moveto +258.0065 230.53394 lineto +253.87564 230.53394 lineto +253.9147 231.15243 254.10025 231.62443 254.43228 231.94995 curveto +254.76757 232.27222 255.23306 232.43335 255.82877 232.43335 curveto +256.17381 232.43335 256.50747 232.39103 256.82974 232.3064 curveto +257.15526 232.22176 257.47752 232.09481 257.79654 231.92554 curveto +257.79654 232.77515 lineto +257.47427 232.91187 257.14387 233.01603 256.80533 233.08765 curveto +256.46678 233.15926 256.12336 233.19507 255.77505 233.19507 curveto +254.90266 233.19507 254.21093 232.94116 253.69986 232.43335 curveto +253.19205 231.92554 252.93814 231.23869 252.93814 230.3728 curveto +252.93814 229.47762 253.17903 228.76799 253.6608 228.2439 curveto +254.14582 227.71656 254.79849 227.45289 255.6188 227.45288 curveto +256.35448 227.45289 256.93553 227.69052 257.36197 228.16577 curveto +257.79165 228.63778 258.00649 229.28068 258.0065 230.09448 curveto +257.10806 229.83081 moveto +257.10155 229.33928 256.9632 228.94703 256.69302 228.65405 curveto +256.42609 228.36109 256.07128 228.2146 255.62857 228.2146 curveto +255.12727 228.2146 254.72525 228.35621 254.42252 228.6394 curveto +254.12303 228.92261 253.95051 229.32137 253.90494 229.83569 curveto +257.10806 229.83081 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +238.41666 242.74585 moveto +238.41666 243.59546 lineto +238.16275 243.46526 237.89907 243.3676 237.62564 243.30249 curveto +237.3522 243.23739 237.069 243.20484 236.77603 243.20483 curveto +236.33007 243.20484 235.99478 243.2732 235.77017 243.40991 curveto +235.54882 243.54664 235.43814 243.75171 235.43814 244.02515 curveto +235.43814 244.23348 235.51789 244.39787 235.6774 244.51831 curveto +235.8369 244.6355 236.15754 244.74781 236.63931 244.85522 curveto +236.94693 244.92358 lineto +237.58495 245.06031 238.03742 245.25399 238.30435 245.50464 curveto +238.57453 245.75204 238.70962 246.09872 238.70963 246.54468 curveto +238.70962 247.05249 238.5078 247.45451 238.10416 247.75073 curveto +237.70376 248.04696 237.152 248.19507 236.44888 248.19507 curveto +236.15591 248.19507 235.84992 248.16577 235.53091 248.10718 curveto +235.21516 248.05184 234.8815 247.9672 234.52994 247.85327 curveto +234.52994 246.92554 lineto +234.86197 247.09806 235.18912 247.22827 235.51138 247.31616 curveto +235.83365 247.4008 236.15266 247.44312 236.46841 247.44312 curveto +236.89159 247.44312 237.21711 247.3715 237.44498 247.22827 curveto +237.67284 247.08179 237.78677 246.87671 237.78677 246.61304 curveto +237.78677 246.3689 237.70376 246.18172 237.53775 246.05151 curveto +237.37499 245.92131 237.01529 245.79598 236.45865 245.67554 curveto +236.14615 245.60229 lineto +235.58951 245.48511 235.18749 245.30607 234.94009 245.06519 curveto +234.6927 244.82105 234.569 244.48739 234.569 244.06421 curveto +234.569 243.54989 234.75129 243.15276 235.11588 242.8728 curveto +235.48046 242.59286 235.99803 242.45289 236.66861 242.45288 curveto +237.00064 242.45289 237.31314 242.4773 237.60611 242.52612 curveto +237.89907 242.57496 238.16926 242.6482 238.41666 242.74585 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +244.69107 244.75269 moveto +244.69107 248.05347 lineto +243.79263 248.05347 lineto +243.79263 244.78198 lineto +243.79263 244.26441 243.69172 243.87704 243.4899 243.61987 curveto +243.28807 243.36272 242.98534 243.23414 242.5817 243.23413 curveto +242.09667 243.23414 241.71418 243.38876 241.43423 243.698 curveto +241.15428 244.00725 241.01431 244.4288 241.01431 244.96265 curveto +241.01431 248.05347 lineto +240.11099 248.05347 lineto +240.11099 240.45581 lineto +241.01431 240.45581 lineto +241.01431 243.43433 lineto +241.22915 243.10556 241.48143 242.85979 241.77115 242.69702 curveto +242.06411 242.53427 242.40103 242.45289 242.78189 242.45288 curveto +243.41014 242.45289 243.8854 242.6482 244.20767 243.03882 curveto +244.52993 243.42619 244.69107 243.99748 244.69107 244.75269 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +248.61197 243.2146 moveto +248.1302 243.2146 247.74934 243.40341 247.46939 243.78101 curveto +247.18944 244.15536 247.04947 244.66968 247.04947 245.32397 curveto +247.04947 245.97827 247.18781 246.49422 247.46451 246.87183 curveto +247.74445 247.24618 248.12694 247.43335 248.61197 247.43335 curveto +249.09048 247.43335 249.46971 247.24455 249.74966 246.86694 curveto +250.02961 246.48934 250.16958 245.97502 250.16959 245.32397 curveto +250.16958 244.67619 250.02961 244.1635 249.74966 243.78589 curveto +249.46971 243.40503 249.09048 243.2146 248.61197 243.2146 curveto +248.61197 242.45288 moveto +249.39322 242.45289 250.00682 242.70679 250.45279 243.2146 curveto +250.89875 243.72242 251.12173 244.42554 251.12173 245.32397 curveto +251.12173 246.21916 250.89875 246.92228 250.45279 247.43335 curveto +250.00682 247.94116 249.39322 248.19507 248.61197 248.19507 curveto +247.82746 248.19507 247.21223 247.94116 246.76627 247.43335 curveto +246.32356 246.92228 246.1022 246.21916 246.1022 245.32397 curveto +246.1022 244.42554 246.32356 243.72242 246.76627 243.2146 curveto +247.21223 242.70679 247.82746 242.45289 248.61197 242.45288 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +252.08365 242.58472 moveto +252.98209 242.58472 lineto +254.10513 246.85229 lineto +255.2233 242.58472 lineto +256.28287 242.58472 lineto +257.40591 246.85229 lineto +258.52408 242.58472 lineto +259.42252 242.58472 lineto +257.99185 248.05347 lineto +256.93228 248.05347 lineto +255.75552 243.57104 lineto +254.57388 248.05347 lineto +253.51431 248.05347 lineto +252.08365 242.58472 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +311.38464 185.46135 moveto +311.38464 188.76213 lineto +310.48621 188.76213 lineto +310.48621 185.49065 lineto +310.4862 184.97307 310.38529 184.5857 310.18347 184.32854 curveto +309.98164 184.07138 309.67891 183.9428 309.27527 183.94279 curveto +308.79024 183.9428 308.40775 184.09742 308.12781 184.40666 curveto +307.84786 184.71591 307.70788 185.13746 307.70789 185.67131 curveto +307.70789 188.76213 lineto +306.80457 188.76213 lineto +306.80457 181.16447 lineto +307.70789 181.16447 lineto +307.70789 184.14299 lineto +307.92273 183.81422 308.17501 183.56845 308.46472 183.40569 curveto +308.75769 183.24293 309.0946 183.16155 309.47546 183.16154 curveto +310.10371 183.16155 310.57897 183.35686 310.90125 183.74748 curveto +311.22351 184.13486 311.38464 184.70615 311.38464 185.46135 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +315.67175 186.01311 moveto +314.94584 186.01311 314.44291 186.09612 314.16296 186.26213 curveto +313.88301 186.42815 313.74304 186.71135 313.74304 187.11174 curveto +313.74304 187.43075 313.84721 187.68466 314.05554 187.87346 curveto +314.26713 188.05901 314.55359 188.15178 314.91492 188.15178 curveto +315.41296 188.15178 315.81172 187.976 316.11121 187.62444 curveto +316.41394 187.26962 316.5653 186.79924 316.56531 186.2133 curveto +316.56531 186.01311 lineto +315.67175 186.01311 lineto +317.46375 185.64201 moveto +317.46375 188.76213 lineto +316.56531 188.76213 lineto +316.56531 187.93205 lineto +316.36023 188.26408 316.10469 188.50985 315.79871 188.66936 curveto +315.49271 188.82561 315.11836 188.90373 314.67566 188.90373 curveto +314.11576 188.90373 313.6698 188.74748 313.33777 188.43498 curveto +313.00899 188.11923 312.8446 187.69768 312.8446 187.17033 curveto +312.8446 186.5551 313.04968 186.09123 313.45984 185.77873 curveto +313.87325 185.46624 314.48848 185.30999 315.30554 185.30998 curveto +316.56531 185.30998 lineto +316.56531 185.22209 lineto +316.5653 184.80868 316.42858 184.48967 316.15515 184.26506 curveto +315.88497 184.0372 315.50411 183.92327 315.01257 183.92326 curveto +314.70007 183.92327 314.39571 183.9607 314.09949 184.03557 curveto +313.80326 184.11044 313.51843 184.22275 313.245 184.37248 curveto +313.245 183.5424 lineto +313.57377 183.41546 313.89278 183.32106 314.20203 183.2592 curveto +314.51127 183.1941 314.81238 183.16155 315.10535 183.16154 curveto +315.89636 183.16155 316.48718 183.36663 316.87781 183.77678 curveto +317.26843 184.18694 317.46374 184.80868 317.46375 185.64201 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +318.67468 183.29338 moveto +319.62683 183.29338 lineto +321.33582 187.88322 lineto +323.0448 183.29338 lineto +323.99695 183.29338 lineto +321.94617 188.76213 lineto +320.72546 188.76213 lineto +318.67468 183.29338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +329.91492 185.80315 moveto +329.91492 186.2426 lineto +325.78406 186.2426 lineto +325.82312 186.86109 326.00867 187.3331 326.3407 187.65862 curveto +326.67598 187.98088 327.14148 188.14201 327.73718 188.14201 curveto +328.08223 188.14201 328.41589 188.0997 328.73816 188.01506 curveto +329.06368 187.93043 329.38594 187.80347 329.70496 187.6342 curveto +329.70496 188.48381 lineto +329.38269 188.62053 329.05228 188.7247 328.71375 188.79631 curveto +328.3752 188.86792 328.03178 188.90373 327.68347 188.90373 curveto +326.81107 188.90373 326.11934 188.64983 325.60828 188.14201 curveto +325.10046 187.6342 324.84656 186.94735 324.84656 186.08147 curveto +324.84656 185.18629 325.08744 184.47665 325.56921 183.95256 curveto +326.05424 183.42522 326.70691 183.16155 327.52722 183.16154 curveto +328.26289 183.16155 328.84395 183.39918 329.27039 183.87444 curveto +329.70007 184.34645 329.91491 184.98935 329.91492 185.80315 curveto +329.01648 185.53947 moveto +329.00996 185.04794 328.87162 184.65569 328.60144 184.36272 curveto +328.33451 184.06975 327.97969 183.92327 327.53699 183.92326 curveto +327.03568 183.92327 326.63366 184.06487 326.33093 184.34807 curveto +326.03145 184.63128 325.85893 185.03004 325.81335 185.54436 curveto +329.01648 185.53947 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +310.32507 198.45451 moveto +310.32507 199.30412 lineto +310.07116 199.17392 309.80749 199.07626 309.53406 199.01115 curveto +309.26062 198.94605 308.97741 198.9135 308.68445 198.9135 curveto +308.23848 198.9135 307.9032 198.98186 307.67859 199.11858 curveto +307.45723 199.2553 307.34656 199.46038 307.34656 199.73381 curveto +307.34656 199.94215 307.42631 200.10654 307.58582 200.22697 curveto +307.74532 200.34417 308.06596 200.45647 308.54773 200.56389 curveto +308.85535 200.63225 lineto +309.49336 200.76897 309.94584 200.96265 310.21277 201.2133 curveto +310.48295 201.4607 310.61804 201.80738 310.61804 202.25334 curveto +310.61804 202.76116 310.41621 203.16317 310.01257 203.4594 curveto +309.61218 203.75562 309.06042 203.90373 308.3573 203.90373 curveto +308.06433 203.90373 307.75834 203.87444 307.43933 203.81584 curveto +307.12357 203.7605 306.78992 203.67587 306.43835 203.56194 curveto +306.43835 202.6342 lineto +306.77038 202.80673 307.09753 202.93694 307.4198 203.02483 curveto +307.74206 203.10946 308.06107 203.15178 308.37683 203.15178 curveto +308.80001 203.15178 309.12553 203.08017 309.35339 202.93694 curveto +309.58125 202.79045 309.69519 202.58537 309.69519 202.3217 curveto +309.69519 202.07756 309.61218 201.89039 309.44617 201.76018 curveto +309.2834 201.62997 308.9237 201.50465 308.36707 201.3842 curveto +308.05457 201.31096 lineto +307.49792 201.19377 307.09591 201.01474 306.84851 200.77385 curveto +306.60111 200.52971 306.47742 200.19605 306.47742 199.77287 curveto +306.47742 199.25855 306.65971 198.86142 307.02429 198.58147 curveto +307.38887 198.30152 307.90645 198.16155 308.57703 198.16154 curveto +308.90905 198.16155 309.22155 198.18596 309.51453 198.23479 curveto +309.80749 198.28362 310.07767 198.35686 310.32507 198.45451 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +316.59949 200.46135 moveto +316.59949 203.76213 lineto +315.70105 203.76213 lineto +315.70105 200.49065 lineto +315.70105 199.97307 315.60013 199.5857 315.39832 199.32854 curveto +315.19649 199.07138 314.89375 198.9428 314.49011 198.94279 curveto +314.00508 198.9428 313.6226 199.09742 313.34265 199.40666 curveto +313.0627 199.71591 312.92273 200.13746 312.92273 200.67131 curveto +312.92273 203.76213 lineto +312.01941 203.76213 lineto +312.01941 196.16447 lineto +312.92273 196.16447 lineto +312.92273 199.14299 lineto +313.13757 198.81422 313.38985 198.56845 313.67957 198.40569 curveto +313.97253 198.24293 314.30945 198.16155 314.69031 198.16154 curveto +315.31856 198.16155 315.79382 198.35686 316.11609 198.74748 curveto +316.43835 199.13486 316.59948 199.70615 316.59949 200.46135 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +320.52039 198.92326 moveto +320.03861 198.92327 319.65775 199.11207 319.37781 199.48967 curveto +319.09786 199.86402 318.95788 200.37835 318.95789 201.03264 curveto +318.95788 201.68694 319.09623 202.20289 319.37292 202.58049 curveto +319.65287 202.95484 320.03536 203.14201 320.52039 203.14201 curveto +320.9989 203.14201 321.37813 202.95321 321.65808 202.57561 curveto +321.93802 202.198 322.078 201.68368 322.078 201.03264 curveto +322.078 200.38486 321.93802 199.87216 321.65808 199.49455 curveto +321.37813 199.1137 320.9989 198.92327 320.52039 198.92326 curveto +320.52039 198.16154 moveto +321.30163 198.16155 321.91524 198.41546 322.36121 198.92326 curveto +322.80716 199.43108 323.03015 200.1342 323.03015 201.03264 curveto +323.03015 201.92782 322.80716 202.63095 322.36121 203.14201 curveto +321.91524 203.64983 321.30163 203.90373 320.52039 203.90373 curveto +319.73588 203.90373 319.12064 203.64983 318.67468 203.14201 curveto +318.23197 202.63095 318.01062 201.92782 318.01062 201.03264 curveto +318.01062 200.1342 318.23197 199.43108 318.67468 198.92326 curveto +319.12064 198.41546 319.73588 198.16155 320.52039 198.16154 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +323.99207 198.29338 moveto +324.8905 198.29338 lineto +326.01355 202.56096 lineto +327.13171 198.29338 lineto +328.19128 198.29338 lineto +329.31433 202.56096 lineto +330.4325 198.29338 lineto +331.33093 198.29338 lineto +329.90027 203.76213 lineto +328.8407 203.76213 lineto +327.66394 199.27971 lineto +326.4823 203.76213 lineto +325.42273 203.76213 lineto +323.99207 198.29338 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +305.63477 140.25864 moveto +305.63477 143.15903 lineto +304.73145 143.15903 lineto +304.73145 135.6102 lineto +305.63477 135.6102 lineto +305.63477 136.44028 lineto +305.82357 136.11476 306.0612 135.87388 306.34766 135.71762 curveto +306.63737 135.55812 306.98242 135.47837 307.38281 135.47836 curveto +308.04687 135.47837 308.58561 135.74204 308.99902 136.26938 curveto +309.41568 136.79673 309.62402 137.49009 309.62402 138.34946 curveto +309.62402 139.20883 309.41568 139.90219 308.99902 140.42953 curveto +308.58561 140.95688 308.04687 141.22055 307.38281 141.22055 curveto +306.98242 141.22055 306.63737 141.14243 306.34766 140.98618 curveto +306.0612 140.82667 305.82357 140.58416 305.63477 140.25864 curveto +308.69141 138.34946 moveto +308.6914 137.68865 308.55468 137.17108 308.28125 136.79672 curveto +308.01106 136.41912 307.63834 136.23032 307.16309 136.23032 curveto +306.68782 136.23032 306.31347 136.41912 306.04004 136.79672 curveto +305.76985 137.17108 305.63476 137.68865 305.63477 138.34946 curveto +305.63476 139.01027 305.76985 139.52947 306.04004 139.90707 curveto +306.31347 140.28142 306.68782 140.4686 307.16309 140.4686 curveto +307.63834 140.4686 308.01106 140.28142 308.28125 139.90707 curveto +308.55468 139.52947 308.6914 139.01027 308.69141 138.34946 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +314.28223 136.45004 moveto +314.18131 136.39145 314.07063 136.34914 313.9502 136.32309 curveto +313.833 136.2938 313.7028 136.27915 313.55957 136.27914 curveto +313.05175 136.27915 312.66113 136.44516 312.3877 136.77719 curveto +312.11751 137.10597 311.98242 137.5796 311.98242 138.19809 curveto +311.98242 141.07895 lineto +311.0791 141.07895 lineto +311.0791 135.6102 lineto +311.98242 135.6102 lineto +311.98242 136.45981 lineto +312.17122 136.12778 312.41699 135.88201 312.71973 135.7225 curveto +313.02246 135.55975 313.3903 135.47837 313.82324 135.47836 curveto +313.88509 135.47837 313.95345 135.48325 314.02832 135.49301 curveto +314.10319 135.49953 314.18619 135.51092 314.27734 135.52719 curveto +314.28223 136.45004 lineto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +317.13867 136.24008 moveto +316.6569 136.24009 316.27604 136.42889 315.99609 136.80649 curveto +315.71614 137.18084 315.57617 137.69516 315.57617 138.34946 curveto +315.57617 139.00376 315.71452 139.51971 315.99121 139.89731 curveto +316.27116 140.27166 316.65364 140.45883 317.13867 140.45883 curveto +317.61718 140.45883 317.99642 140.27003 318.27637 139.89243 curveto +318.55631 139.51482 318.69628 139.0005 318.69629 138.34946 curveto +318.69628 137.70167 318.55631 137.18898 318.27637 136.81137 curveto +317.99642 136.43052 317.61718 136.24009 317.13867 136.24008 curveto +317.13867 135.47836 moveto +317.91992 135.47837 318.53352 135.73227 318.97949 136.24008 curveto +319.42545 136.7479 319.64843 137.45102 319.64844 138.34946 curveto +319.64843 139.24464 319.42545 139.94777 318.97949 140.45883 curveto +318.53352 140.96664 317.91992 141.22055 317.13867 141.22055 curveto +316.35416 141.22055 315.73893 140.96664 315.29297 140.45883 curveto +314.85026 139.94777 314.62891 139.24464 314.62891 138.34946 curveto +314.62891 137.45102 314.85026 136.7479 315.29297 136.24008 curveto +315.73893 135.73227 316.35416 135.47837 317.13867 135.47836 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +323.25195 136.24008 moveto +322.77018 136.24009 322.38932 136.42889 322.10938 136.80649 curveto +321.82943 137.18084 321.68945 137.69516 321.68945 138.34946 curveto +321.68945 139.00376 321.8278 139.51971 322.10449 139.89731 curveto +322.38444 140.27166 322.76692 140.45883 323.25195 140.45883 curveto +323.73047 140.45883 324.1097 140.27003 324.38965 139.89243 curveto +324.66959 139.51482 324.80957 139.0005 324.80957 138.34946 curveto +324.80957 137.70167 324.66959 137.18898 324.38965 136.81137 curveto +324.1097 136.43052 323.73047 136.24009 323.25195 136.24008 curveto +323.25195 135.47836 moveto +324.0332 135.47837 324.64681 135.73227 325.09277 136.24008 curveto +325.53873 136.7479 325.76171 137.45102 325.76172 138.34946 curveto +325.76171 139.24464 325.53873 139.94777 325.09277 140.45883 curveto +324.64681 140.96664 324.0332 141.22055 323.25195 141.22055 curveto +322.46745 141.22055 321.85221 140.96664 321.40625 140.45883 curveto +320.96354 139.94777 320.74219 139.24464 320.74219 138.34946 curveto +320.74219 137.45102 320.96354 136.7479 321.40625 136.24008 curveto +321.85221 135.73227 322.46745 135.47837 323.25195 135.47836 curveto +fill +grestore +gsave +0 0 0 setrgbcolor +newpath +330.01465 133.48129 moveto +330.01465 134.22836 lineto +329.15527 134.22836 lineto +328.83301 134.22837 328.6084 134.29347 328.48145 134.42368 curveto +328.35775 134.55389 328.2959 134.78827 328.2959 135.1268 curveto +328.2959 135.6102 lineto +329.77539 135.6102 lineto +329.77539 136.30844 lineto +328.2959 136.30844 lineto +328.2959 141.07895 lineto +327.39258 141.07895 lineto +327.39258 136.30844 lineto +326.5332 136.30844 lineto +326.5332 135.6102 lineto +327.39258 135.6102 lineto +327.39258 135.22934 lineto +327.39258 134.62062 327.53418 134.17791 327.81738 133.90121 curveto +328.10058 133.62127 328.5498 133.4813 329.16504 133.48129 curveto +330.01465 133.48129 lineto +fill +grestore +grestore +grestore +showpage +%%EOF diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/isar-vm.pdf Binary file doc-src/IsarRef/document/isar-vm.pdf has changed diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/isar-vm.svg --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/document/isar-vm.svg Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,460 @@ + + + + + + + + + + + + + + + + + + + + + + + + + image/svg+xml + + + + + + + + chain + + + + + + + + prove + + + + + + state + + + + + + + + theorem + + qed + qed + fixassume + { }next + notelet + usingunfolding + then + haveshow + haveshow + proof + + diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/document/root.tex Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,91 @@ +\documentclass[12pt,a4paper,fleqn]{report} +\usepackage{amssymb} +\usepackage{eurosym} +\usepackage[english]{babel} +\usepackage[only,bigsqcap]{stmaryrd} +\usepackage{textcomp} +\usepackage{latexsym} +\usepackage{graphicx} +\let\intorig=\int %iman.sty redefines \int +\usepackage{iman,extra,isar,proof} +\usepackage[nohyphen,strings]{underscore} +\usepackage{isabelle} +\usepackage{isabellesym} +\usepackage{railsetup} +\usepackage{ttbox} +\usepackage{supertabular} +\usepackage{style} +\usepackage{pdfsetup} + +\hyphenation{Isabelle} +\hyphenation{Isar} + +\isadroptag{theory} +\title{\includegraphics[scale=0.5]{isabelle_isar} \\[4ex] The Isabelle/Isar Reference Manual} +\author{\emph{Makarius Wenzel} \\[3ex] + With Contributions by + Clemens Ballarin, + Stefan Berghofer, \\ + Jasmin Blanchette, + Timothy Bourke, + Lukas Bulwahn, \\ + Lucas Dixon, + Florian Haftmann, + Brian Huffman, \\ + Gerwin Klein, + Alexander Krauss, + Ond\v{r}ej Kun\v{c}ar, \\ + Tobias Nipkow, + Lars Noschinski, + David von Oheimb, \\ + Larry Paulson, + Sebastian Skalberg +} + +\makeindex + +\chardef\charbackquote=`\` +\newcommand{\backquote}{\mbox{\tt\charbackquote}} + + +\begin{document} + +\maketitle + +\pagenumbering{roman} +{\def\isamarkupchapter#1{\chapter*{#1}}\input{Preface.tex}} +\tableofcontents +\clearfirst + +\part{Basic Concepts} +\input{Synopsis.tex} +\input{Framework.tex} +\input{First_Order_Logic.tex} +\part{General Language Elements} +\input{Outer_Syntax.tex} +\input{Document_Preparation.tex} +\input{Spec.tex} +\input{Proof.tex} +\input{Inner_Syntax.tex} +\input{Misc.tex} +\input{Generic.tex} +\part{Object-Logic} +\input{HOL_Specific.tex} + +\part{Appendix} +\appendix +\input{Quick_Reference.tex} +\let\int\intorig +\input{Symbols.tex} +\input{ML_Tactic.tex} + +\begingroup + \tocentry{\bibname} + \bibliographystyle{abbrv} \small\raggedright\frenchspacing + \bibliography{manual} +\endgroup + +\tocentry{\indexname} +\printindex + +\end{document} diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/showsymbols --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/document/showsymbols Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,26 @@ +#!/usr/bin/env perl + +print "\\begin{supertabular}{ll\@{\\qquad}ll}\n"; + +$eol = "&"; + +while () { + if (m/^\\newcommand\{\\isasym([A-Za-z]+)\}/) { + print "\\verb,\\<$1>, & {\\isasym$1} $eol\n"; +# print "\\verb,\\<$1>, & \\isactrlbold{\\isasym$1}~{\\isasym$1} $eol\n"; +# print "\\verb,\\<$1>, & B\\isactrlsup{\\isasym$1} $eol\n"; +# print "\\verb,\\<$1>, & B\\isactrlsub{\\isasym$1} $eol\n"; + if ("$eol" eq "&") { + $eol = "\\\\"; + } else { + $eol = "&"; + } + } +} + +if ("$eol" eq "\\\\") { + print "$eol\n"; +} + +print "\\end{supertabular}\n"; + diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/document/style.sty --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarRef/document/style.sty Tue Aug 28 12:45:49 2012 +0200 @@ -0,0 +1,47 @@ +%% toc +\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} +\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} + +\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} + +%% references +\newcommand{\secref}[1]{\S\ref{#1}} +\newcommand{\chref}[1]{chapter~\ref{#1}} +\newcommand{\Chref}[1]{Chapter~\ref{#1}} +\newcommand{\appref}[1]{appendix~\ref{#1}} +\newcommand{\Appref}[1]{Appendix~\ref{#1}} +\newcommand{\figref}[1]{figure~\ref{#1}} +\newcommand{\Figref}[1]{Figure~\ref{#1}} + +%% Isar +\newcommand{\isasymBBAR}{{\,\newdimen{\tmpheight}\settoheight\tmpheight{\isacharbar}\rule{1pt}{\tmpheight}\,}} +\isafoldtag{noproof}\def\isafoldnoproof{~\isafold{proof}} +\newcommand{\isadigitreset}{\def\isadigit##1{##1}} +\renewcommand{\isacommand}[1]{\isakeyword{\isadigitreset#1}} + +%% ML +\newenvironment{mldecls}{\par\noindent\begingroup\def\isanewline{\\}\begin{tabular}{ll}}{\end{tabular}\medskip\endgroup} + +\renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\isadigitreset} +\renewcommand{\endisatagML}{\endgroup} + +%% math +\newcommand{\isasymstrut}{\isamath{\mathstrut}} +\newcommand{\isasymvartheta}{\isamath{\,\theta}} +\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} +\renewcommand{\isadigit}[1]{\isamath{#1}} +\newcommand{\text}[1]{\mbox{#1}} + +%% global style options +\pagestyle{headings} +\sloppy + +\parindent 0pt\parskip 0.5ex + +\isabellestyle{literal} + +\newcommand{\isasymdash}{\isatext{\mbox{-}}} + +\railtermfont{\isabellestyle{tt}} +\railnontermfont{\isabellestyle{literal}} +\railnamefont{\isabellestyle{literal}} diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/isar-ref.tex --- a/doc-src/IsarRef/isar-ref.tex Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,91 +0,0 @@ -\documentclass[12pt,a4paper,fleqn]{report} -\usepackage{amssymb} -\usepackage{eurosym} -\usepackage[english]{babel} -\usepackage[only,bigsqcap]{stmaryrd} -\usepackage{textcomp} -\usepackage{latexsym} -\usepackage{graphicx} -\let\intorig=\int %iman.sty redefines \int -\usepackage{../iman,../extra,../isar,../proof} -\usepackage[nohyphen,strings]{../underscore} -\usepackage{../../lib/texinputs/isabelle} -\usepackage{../../lib/texinputs/isabellesym} -\usepackage{../../lib/texinputs/railsetup} -\usepackage{../ttbox} -\usepackage{supertabular} -\usepackage{style} -\usepackage{../pdfsetup} - -\hyphenation{Isabelle} -\hyphenation{Isar} - -\isadroptag{theory} -\title{\includegraphics[scale=0.5]{isabelle_isar} \\[4ex] The Isabelle/Isar Reference Manual} -\author{\emph{Makarius Wenzel} \\[3ex] - With Contributions by - Clemens Ballarin, - Stefan Berghofer, \\ - Jasmin Blanchette, - Timothy Bourke, - Lukas Bulwahn, \\ - Lucas Dixon, - Florian Haftmann, - Brian Huffman, \\ - Gerwin Klein, - Alexander Krauss, - Ond\v{r}ej Kun\v{c}ar, \\ - Tobias Nipkow, - Lars Noschinski, - David von Oheimb, \\ - Larry Paulson, - Sebastian Skalberg -} - -\makeindex - -\chardef\charbackquote=`\` -\newcommand{\backquote}{\mbox{\tt\charbackquote}} - - -\begin{document} - -\maketitle - -\pagenumbering{roman} -{\def\isamarkupchapter#1{\chapter*{#1}}\input{Thy/document/Preface.tex}} -\tableofcontents -\clearfirst - -\part{Basic Concepts} -\input{Thy/document/Synopsis.tex} -\input{Thy/document/Framework.tex} -\input{Thy/document/First_Order_Logic.tex} -\part{General Language Elements} -\input{Thy/document/Outer_Syntax.tex} -\input{Thy/document/Document_Preparation.tex} -\input{Thy/document/Spec.tex} -\input{Thy/document/Proof.tex} -\input{Thy/document/Inner_Syntax.tex} -\input{Thy/document/Misc.tex} -\input{Thy/document/Generic.tex} -\part{Object-Logic} -\input{Thy/document/HOL_Specific.tex} - -\part{Appendix} -\appendix -\input{Thy/document/Quick_Reference.tex} -\let\int\intorig -\input{Thy/document/Symbols.tex} -\input{Thy/document/ML_Tactic.tex} - -\begingroup - \tocentry{\bibname} - \bibliographystyle{abbrv} \small\raggedright\frenchspacing - \bibliography{../manual} -\endgroup - -\tocentry{\indexname} -\printindex - -\end{document} diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/showsymbols --- a/doc-src/IsarRef/showsymbols Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,26 +0,0 @@ -#!/usr/bin/env perl - -print "\\begin{supertabular}{ll\@{\\qquad}ll}\n"; - -$eol = "&"; - -while () { - if (m/^\\newcommand\{\\isasym([A-Za-z]+)\}/) { - print "\\verb,\\<$1>, & {\\isasym$1} $eol\n"; -# print "\\verb,\\<$1>, & \\isactrlbold{\\isasym$1}~{\\isasym$1} $eol\n"; -# print "\\verb,\\<$1>, & B\\isactrlsup{\\isasym$1} $eol\n"; -# print "\\verb,\\<$1>, & B\\isactrlsub{\\isasym$1} $eol\n"; - if ("$eol" eq "&") { - $eol = "\\\\"; - } else { - $eol = "&"; - } - } -} - -if ("$eol" eq "\\\\") { - print "$eol\n"; -} - -print "\\end{supertabular}\n"; - diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/IsarRef/style.sty --- a/doc-src/IsarRef/style.sty Tue Aug 28 12:31:53 2012 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,47 +0,0 @@ -%% toc -\newcommand{\tocentry}[1]{\cleardoublepage\phantomsection\addcontentsline{toc}{chapter}{#1} -\@mkboth{\MakeUppercase{#1}}{\MakeUppercase{#1}}} - -\setcounter{secnumdepth}{2} \setcounter{tocdepth}{2} - -%% references -\newcommand{\secref}[1]{\S\ref{#1}} -\newcommand{\chref}[1]{chapter~\ref{#1}} -\newcommand{\Chref}[1]{Chapter~\ref{#1}} -\newcommand{\appref}[1]{appendix~\ref{#1}} -\newcommand{\Appref}[1]{Appendix~\ref{#1}} -\newcommand{\figref}[1]{figure~\ref{#1}} -\newcommand{\Figref}[1]{Figure~\ref{#1}} - -%% Isar -\newcommand{\isasymBBAR}{{\,\newdimen{\tmpheight}\settoheight\tmpheight{\isacharbar}\rule{1pt}{\tmpheight}\,}} -\isafoldtag{noproof}\def\isafoldnoproof{~\isafold{proof}} -\newcommand{\isadigitreset}{\def\isadigit##1{##1}} -\renewcommand{\isacommand}[1]{\isakeyword{\isadigitreset#1}} - -%% ML -\newenvironment{mldecls}{\par\noindent\begingroup\def\isanewline{\\}\begin{tabular}{ll}}{\end{tabular}\medskip\endgroup} - -\renewcommand{\isatagML}{\begingroup\isabellestyle{default}\isastyle\isadigitreset} -\renewcommand{\endisatagML}{\endgroup} - -%% math -\newcommand{\isasymstrut}{\isamath{\mathstrut}} -\newcommand{\isasymvartheta}{\isamath{\,\theta}} -\newcommand{\isactrlvec}[1]{\emph{$\overline{#1}$}} -\renewcommand{\isadigit}[1]{\isamath{#1}} -\newcommand{\text}[1]{\mbox{#1}} - -%% global style options -\pagestyle{headings} -\sloppy - -\parindent 0pt\parskip 0.5ex - -\isabellestyle{literal} - -\newcommand{\isasymdash}{\isatext{\mbox{-}}} - -\railtermfont{\isabellestyle{tt}} -\railnontermfont{\isabellestyle{literal}} -\railnamefont{\isabellestyle{literal}} diff -r c04001b3a753 -r 12afbf6eb7f9 doc-src/ROOT --- a/doc-src/ROOT Tue Aug 28 12:31:53 2012 +0200 +++ b/doc-src/ROOT Tue Aug 28 12:45:49 2012 +0200 @@ -82,10 +82,8 @@ "document/build" "document/root.tex" -session IsarRef (doc) in "IsarRef/Thy" = HOL + - options [browser_info = false, document = false, - document_dump = document, document_dump_mode = "tex", - quick_and_dirty, thy_output_source] +session IsarRef (doc) in "IsarRef" = HOL + + options [document_variants = "isar-ref", quick_and_dirty, thy_output_source] theories Preface Synopsis @@ -102,6 +100,21 @@ Quick_Reference Symbols ML_Tactic + files + "../pdfsetup.sty" + "../iman.sty" + "../extra.sty" + "../ttbox.sty" + "../proof.sty" + "../isar.sty" + "../manual.bib" + "document/build" + "document/isar-vm.eps" + "document/isar-vm.pdf" + "document/isar-vm.svg" + "document/root.tex" + "document/showsymbols" + "document/style.sty" session LaTeXsugar (doc) in "LaTeXsugar" = HOL + options [document_variants = "sugar"]