# HG changeset patch # User wenzelm # Date 1329219655 -3600 # Node ID 915af80f74b380e723c9716ef9b6a423d96b60fb # Parent a1c7b842ff8be6547a350cac3c097dca353a5c97# Parent ec2e20b27638494bcf246776ac4c8e229e6270df merged, resolving trivial conflicts; diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/IsaMakefile --- a/doc-src/IsarImplementation/IsaMakefile Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/IsaMakefile Tue Feb 14 12:40:55 2012 +0100 @@ -21,10 +21,10 @@ Thy: $(LOG)/HOL-Thy.gz -$(LOG)/HOL-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Integration.thy \ - Thy/Isar.thy Thy/Local_Theory.thy Thy/Logic.thy Thy/Prelim.thy \ - Thy/Proof.thy Thy/Syntax.thy Thy/Tactic.thy Thy/ML.thy \ - ../antiquote_setup.ML +$(LOG)/HOL-Thy.gz: Thy/ROOT.ML Thy/Base.thy Thy/Eq.thy \ + Thy/Integration.thy Thy/Isar.thy Thy/Local_Theory.thy Thy/Logic.thy \ + Thy/Prelim.thy Thy/Proof.thy Thy/Syntax.thy Thy/Tactic.thy \ + Thy/ML.thy ../antiquote_setup.ML @$(USEDIR) HOL Thy @rm -f Thy/document/isabelle.sty Thy/document/isabellesym.sty \ Thy/document/pdfsetup.sty Thy/document/session.tex diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Makefile --- a/doc-src/IsarImplementation/Makefile Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Makefile Tue Feb 14 12:40:55 2012 +0100 @@ -13,11 +13,12 @@ FILES = ../extra.sty ../iman.sty ../../lib/texinputs/isabelle.sty \ ../../lib/texinputs/isabellesym.sty \ ../../lib/texinputs/railsetup.sty ../isar.sty ../manual.bib \ - ../pdfsetup.sty ../proof.sty Thy/document/Integration.tex \ - Thy/document/Isar.tex Thy/document/Local_Theory.tex \ - Thy/document/Logic.tex Thy/document/Prelim.tex \ - Thy/document/Proof.tex Thy/document/Syntax.tex \ - Thy/document/Tactic.tex implementation.tex style.sty + ../pdfsetup.sty ../proof.sty Thy/document/Eq.tex \ + Thy/document/Integration.tex Thy/document/Isar.tex \ + Thy/document/Local_Theory.tex Thy/document/Logic.tex \ + Thy/document/Prelim.tex Thy/document/Proof.tex \ + Thy/document/Syntax.tex Thy/document/Tactic.tex implementation.tex \ + style.sty dvi: $(NAME).dvi diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/Eq.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/Eq.thy Tue Feb 14 12:40:55 2012 +0100 @@ -0,0 +1,75 @@ +theory Eq +imports Base +begin + +chapter {* Equational reasoning *} + +text {* Equality is one of the most fundamental concepts of + mathematics. The Isabelle/Pure logic (\chref{ch:logic}) provides a + builtin relation @{text "\ :: \ \ \ \ prop"} that expresses equality + of arbitrary terms (or propositions) at the framework level, as + expressed by certain basic inference rules (\secref{sec:eq-rules}). + + Equational reasoning means to replace equals by equals, using + reflexivity and transitivity to form chains of replacement steps, + and congruence rules to access sub-structures. Conversions + (\secref{sec:conv}) provide a convenient framework to compose basic + equational steps to build specific equational reasoning tools. + + Higher-order matching is able to provide suitable instantiations for + giving equality rules, which leads to the versatile concept of + @{text "\"}-term rewriting (\secref{sec:rewriting}). Internally + this is based on the general-purpose Simplifier engine of Isabelle, + which is more specific and more efficient than plain conversions. + + Object-logics usually introduce specific notions of equality or + equivalence, and relate it with the Pure equality. This enables to + re-use the Pure tools for equational reasoning for particular + object-logic connectives as well. +*} + + +section {* Basic equality rules \label{sec:eq-rules} *} + +text {* FIXME *} + + +section {* Conversions \label{sec:conv} *} + +text {* FIXME *} + + +section {* Rewriting \label{sec:rewriting} *} + +text {* Rewriting normalizes a given term (theorem or goal) by + replacing instances of given equalities @{text "t \ u"} in subterms. + Rewriting continues until no rewrites are applicable to any subterm. + This may be used to unfold simple definitions of the form @{text "f + x\<^sub>1 \ x\<^sub>n \ u"}, but is slightly more general than that. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML rewrite_goal_tac: "thm list -> int -> tactic"} \\ + @{index_ML rewrite_goals_tac: "thm list -> tactic"} \\ + @{index_ML fold_goals_tac: "thm list -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML rewrite_goal_tac}~@{text "rules i"} rewrites subgoal + @{text "i"} by the given rewrite rules. + + \item @{ML rewrite_goals_tac}~@{text "rules"} rewrites all subgoals + by the given rewrite rules. + + \item @{ML fold_goals_tac}~@{text "rules"} essentially uses @{ML + rewrite_goals_tac} with the symmetric form of each member of @{text + "rules"}, re-ordered to fold longer expression first. This supports + to idea to fold primitive definitions that appear in expended form + in the proof state. + + \end{description} +*} + +end diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/Isar.thy --- a/doc-src/IsarImplementation/Thy/Isar.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/Isar.thy Tue Feb 14 12:40:55 2012 +0100 @@ -311,7 +311,6 @@ @{index_ML METHOD: "(thm list -> tactic) -> Proof.method"} \\ @{index_ML SIMPLE_METHOD: "tactic -> Proof.method"} \\ @{index_ML SIMPLE_METHOD': "(int -> tactic) -> Proof.method"} \\ - @{index_ML HEADGOAL: "(int -> tactic) -> tactic"} \\ @{index_ML Method.insert_tac: "thm list -> int -> tactic"} \\ @{index_ML Method.setup: "binding -> (Proof.context -> Proof.method) context_parser -> string -> theory -> theory"} \\ @@ -336,13 +335,9 @@ applied. \item @{ML SIMPLE_METHOD'}~@{text "tactic"} wraps a tactic that - addresses a specific subgoal as simple proof method. Goal facts are - already inserted into the first subgoal before @{text "tactic"} is - applied to the same. - - \item @{ML HEADGOAL}~@{text "tactic"} applies @{text "tactic"} to - the first subgoal. This is convenient to reproduce part the @{ML - SIMPLE_METHOD'} wrapping within regular @{ML METHOD}, for example. + addresses a specific subgoal as simple proof method that operates on + subgoal 1. Goal facts are inserted into the subgoal then the @{text + "tactic"} is applied. \item @{ML Method.insert_tac}~@{text "facts i"} inserts @{text "facts"} into subgoal @{text "i"}. This is convenient to reproduce diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/Logic.thy --- a/doc-src/IsarImplementation/Thy/Logic.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/Logic.thy Tue Feb 14 12:40:55 2012 +0100 @@ -356,7 +356,7 @@ text %mlref {* \begin{mldecls} @{index_ML_type term} \\ - @{index_ML "op aconv": "term * term -> bool"} \\ + @{index_ML_op "aconv": "term * term -> bool"} \\ @{index_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{index_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{index_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @@ -380,7 +380,7 @@ \item Type @{ML_type term} represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{ML Bound}, @{ML Free}, @{ML - Var}, @{ML Const}, @{ML Abs}, @{ML "op $"}. + Var}, @{ML Const}, @{ML Abs}, @{ML_op "$"}. \item @{text "t"}~@{ML_text aconv}~@{text "u"} checks @{text "\"}-equivalence of two terms. This is the basic equality relation @@ -633,10 +633,18 @@ text %mlref {* \begin{mldecls} + @{index_ML Logic.all: "term -> term -> term"} \\ + @{index_ML Logic.mk_implies: "term * term -> term"} \\ + \end{mldecls} + \begin{mldecls} @{index_ML_type ctyp} \\ @{index_ML_type cterm} \\ @{index_ML Thm.ctyp_of: "theory -> typ -> ctyp"} \\ @{index_ML Thm.cterm_of: "theory -> term -> cterm"} \\ + @{index_ML Thm.capply: "cterm -> cterm -> cterm"} \\ + @{index_ML Thm.cabs: "cterm -> cterm -> cterm"} \\ + @{index_ML Thm.all: "cterm -> cterm -> cterm"} \\ + @{index_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{index_ML_type thm} \\ @@ -663,20 +671,40 @@ \begin{description} + \item @{ML Logic.all}~@{text "a B"} produces a Pure quantification + @{text "\a. B"}, where occurrences of the atomic term @{text "a"} in + the body proposition @{text "B"} are replaced by bound variables. + (See also @{ML lambda} on terms.) + + \item @{ML Logic.mk_implies}~@{text "(A, B)"} produces a Pure + implication @{text "A \ B"}. + \item Types @{ML_type ctyp} and @{ML_type cterm} represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type - constructors, constants etc. in the theory. + constructors, constants etc.\ in the background theory. The + abstract types @{ML_type ctyp} and @{ML_type cterm} are part of the + same inference kernel that is mainly responsible for @{ML_type thm}. + Thus syntactic operations on @{ML_type ctyp} and @{ML_type cterm} + are located in the @{ML_struct Thm} module, even though theorems are + not yet involved at that stage. \item @{ML Thm.ctyp_of}~@{text "thy \"} and @{ML Thm.cterm_of}~@{text "thy t"} explicitly checks types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the theory context. + Full re-certification is relatively slow and should be avoided in + tight reasoning loops. - Re-certification is relatively slow and should be avoided in tight - reasoning loops. There are separate operations to decompose - certified entities (including actual theorems). + \item @{ML Thm.capply}, @{ML Thm.cabs}, @{ML Thm.all}, @{ML + Drule.mk_implies} etc.\ compose certified terms (or propositions) + incrementally. This is equivalent to @{ML Thm.cterm_of} after + unchecked @{ML_op "$"}, @{ML lambda}, @{ML Logic.all}, @{ML + Logic.mk_implies} etc., but there can be a big difference in + performance when large existing entities are composed by a few extra + constructions on top. There are separate operations to decompose + certified terms and theorems to produce certified terms again. \item Type @{ML_type thm} represents proven propositions. This is an abstract datatype that guarantees that its values have been @@ -802,12 +830,13 @@ *} -subsection {* Auxiliary definitions \label{sec:logic-aux} *} +subsection {* Auxiliary connectives \label{sec:logic-aux} *} -text {* - Theory @{text "Pure"} provides a few auxiliary definitions, see - \figref{fig:pure-aux}. These special constants are normally not - exposed to the user, but appear in internal encodings. +text {* Theory @{text "Pure"} provides a few auxiliary connectives + that are defined on top of the primitive ones, see + \figref{fig:pure-aux}. These special constants are useful in + certain internal encodings, and are normally not directly exposed to + the user. \begin{figure}[htb] \begin{center} @@ -1055,23 +1084,50 @@ text %mlref {* \begin{mldecls} - @{index_ML "op RS": "thm * thm -> thm"} \\ - @{index_ML "op OF": "thm * thm list -> thm"} \\ + @{index_ML_op "RSN": "thm * (int * thm) -> thm"} \\ + @{index_ML_op "RS": "thm * thm -> thm"} \\ + + @{index_ML_op "RLN": "thm list * (int * thm list) -> thm list"} \\ + @{index_ML_op "RL": "thm list * thm list -> thm list"} \\ + + @{index_ML_op "MRS": "thm list * thm -> thm"} \\ + @{index_ML_op "OF": "thm * thm list -> thm"} \\ \end{mldecls} \begin{description} - \item @{text "rule\<^sub>1 RS rule\<^sub>2"} resolves @{text "rule\<^sub>1"} with @{text - "rule\<^sub>2"} according to the @{inference resolution} principle - explained above. Note that the corresponding rule attribute in the - Isar language is called @{attribute THEN}. + \item @{text "rule\<^sub>1 RSN (i, rule\<^sub>2)"} resolves the conclusion of + @{text "rule\<^sub>1"} with the @{text i}-th premise of @{text "rule\<^sub>2"}, + according to the @{inference resolution} principle explained above. + Unless there is precisely one resolvent it raises exception @{ML + THM}. + + This corresponds to the rule attribute @{attribute THEN} in Isar + source language. + + \item @{text "rule\<^sub>1 RS rule\<^sub>2"} abbreviates @{text "rule\<^sub>1 RS (1, + rule\<^sub>2)"}. - \item @{text "rule OF rules"} resolves a list of rules with the - first rule, addressing its premises @{text "1, \, length rules"} - (operating from last to first). This means the newly emerging - premises are all concatenated, without interfering. Also note that - compared to @{text "RS"}, the rule argument order is swapped: @{text - "rule\<^sub>1 RS rule\<^sub>2 = rule\<^sub>2 OF [rule\<^sub>1]"}. + \item @{text "rules\<^sub>1 RLN (i, rules\<^sub>2)"} joins lists of rules. For + every @{text "rule\<^sub>1"} in @{text "rules\<^sub>1"} and @{text "rule\<^sub>2"} in + @{text "rules\<^sub>2"}, it resolves the conclusion of @{text "rule\<^sub>1"} with + the @{text "i"}-th premise of @{text "rule\<^sub>2"}, accumulating multiple + results in one big list. Note that such strict enumerations of + higher-order unifications can be inefficient compared to the lazy + variant seen in elementary tactics like @{ML resolve_tac}. + + \item @{text "rules\<^sub>1 RL rules\<^sub>2"} abbreviates @{text "rules\<^sub>1 RLN (1, + rules\<^sub>2)"}. + + \item @{text "[rule\<^sub>1, \, rule\<^sub>n] MRS rule"} resolves @{text "rule\<^isub>i"} + against premise @{text "i"} of @{text "rule"}, for @{text "i = n, \, + 1"}. By working from right to left, newly emerging premises are + concatenated in the result, without interfering. + + \item @{text "rule OF rules"} abbreviates @{text "rules MRS rule"}. + + This corresponds to the rule attribute @{attribute OF} in Isar + source language. \end{description} *} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/ML.thy --- a/doc-src/IsarImplementation/Thy/ML.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/ML.thy Tue Feb 14 12:40:55 2012 +0100 @@ -849,10 +849,10 @@ text %mlref {* \begin{mldecls} - @{index_ML "op |> ": "'a * ('a -> 'b) -> 'b"} \\ - @{index_ML "op |-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ - @{index_ML "op #> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ - @{index_ML "op #-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ + @{index_ML_op "|> ": "'a * ('a -> 'b) -> 'b"} \\ + @{index_ML_op "|-> ": "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ + @{index_ML_op "#> ": "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ + @{index_ML_op "#-> ": "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ \end{mldecls} %FIXME description!? @@ -1465,7 +1465,7 @@ @{index_ML_type "'a Unsynchronized.ref"} \\ @{index_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\ @{index_ML "!": "'a Unsynchronized.ref -> 'a"} \\ - @{index_ML "op :=": "'a Unsynchronized.ref * 'a -> unit"} \\ + @{index_ML_op ":=": "'a Unsynchronized.ref * 'a -> unit"} \\ \end{mldecls} *} @@ -1477,7 +1477,7 @@ The unwieldy name of @{ML Unsynchronized.ref} for the constructor for references in Isabelle/ML emphasizes the inconveniences caused by - mutability. Existing operations @{ML "!"} and @{ML "op :="} are + mutability. Existing operations @{ML "!"} and @{ML_op ":="} are unchanged, but should be used with special precautions, say in a strictly local situation that is guaranteed to be restricted to sequential evaluation --- now and in the future. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/Proof.thy --- a/doc-src/IsarImplementation/Thy/Proof.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/Proof.thy Tue Feb 14 12:40:55 2012 +0100 @@ -387,6 +387,7 @@ text %mlref {* \begin{mldecls} + @{index_ML SELECT_GOAL: "tactic -> int -> tactic"} \\ @{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> @@ -413,6 +414,10 @@ \begin{description} + \item @{ML SELECT_GOAL}~@{text "tac i"} confines a tactic to the + specified subgoal @{text "i"}. This introduces a nested goal state, + without decomposing the internal structure of the subgoal yet. + \item @{ML SUBPROOF}~@{text "tac ctxt i"} decomposes the structure of the specified sub-goal, producing an extended context and a reduced goal, which needs to be solved by the given tactic. All diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/ROOT.ML --- a/doc-src/IsarImplementation/Thy/ROOT.ML Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/ROOT.ML Tue Feb 14 12:40:55 2012 +0100 @@ -1,4 +1,5 @@ use_thys [ + "Eq", "Integration", "Isar", "Local_Theory", diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/Tactic.thy --- a/doc-src/IsarImplementation/Thy/Tactic.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/Tactic.thy Tue Feb 14 12:40:55 2012 +0100 @@ -164,7 +164,7 @@ explicitly, and violating them merely results in ill-behaved tactics experienced by the user (e.g.\ tactics that insist in being applicable only to singleton goals, or prevent composition via - standard tacticals). + standard tacticals such as @{ML REPEAT}). *} text %mlref {* @@ -282,7 +282,11 @@ premises. \item @{ML eresolve_tac}~@{text "thms i"} performs elim-resolution - with the given theorems, which should normally be elimination rules. + with the given theorems, which are normally be elimination rules. + + Note that @{ML "eresolve_tac [asm_rl]"} is equivalent to @{ML + assume_tac}, which facilitates mixing of assumption steps with + genuine eliminations. \item @{ML dresolve_tac}~@{text "thms i"} performs destruct-resolution with the given theorems, which should normally @@ -364,7 +368,9 @@ @{index_ML res_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ @{index_ML eres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ @{index_ML dres_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ - @{index_ML forw_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\[1ex] + @{index_ML forw_inst_tac: "Proof.context -> (indexname * string) list -> thm -> int -> tactic"} \\ + @{index_ML subgoal_tac: "Proof.context -> string -> int -> tactic"} \\ + @{index_ML thin_tac: "Proof.context -> string -> int -> tactic"} \\ @{index_ML rename_tac: "string list -> int -> tactic"} \\ \end{mldecls} @@ -383,6 +389,17 @@ \item @{ML forw_inst_tac} is like @{ML dres_inst_tac} except that the selected assumption is not deleted. + \item @{ML subgoal_tac}~@{text "ctxt \ i"} adds the proposition + @{text "\"} as local premise to subgoal @{text "i"}, and poses the + same as a new subgoal @{text "i + 1"} (in the original context). + + \item @{ML thin_tac}~@{text "ctxt \ i"} deletes the specified + premise from subgoal @{text i}. 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 @{ML rename_tac}~@{text "names i"} renames the innermost parameters of subgoal @{text i} according to the provided @{text names} (which need to be distinct indentifiers). @@ -397,17 +414,448 @@ *} +subsection {* Rearranging goal states *} + +text {* In rare situations there is a need to rearrange goal states: + either the overall collection of subgoals, or the local structure of + a subgoal. Various administrative tactics allow to operate on the + concrete presentation these conceptual sets of formulae. *} + +text %mlref {* + \begin{mldecls} + @{index_ML rotate_tac: "int -> int -> tactic"} \\ + @{index_ML distinct_subgoals_tac: tactic} \\ + @{index_ML flexflex_tac: tactic} \\ + \end{mldecls} + + \begin{description} + + \item @{ML rotate_tac}~@{text "n i"} rotates the premises of subgoal + @{text i} by @{text n} positions: from right to left if @{text n} is + positive, and from left to right if @{text n} is negative. + + \item @{ML distinct_subgoals_tac} removes duplicate subgoals from a + proof state. This is potentially inefficient. + + \item @{ML flexflex_tac} removes all flex-flex pairs from the proof + state by applying the trivial unifier. This drastic step loses + information. It is already part of the Isar infrastructure for + facts resulting from goals, and rarely needs to be invoked manually. + + Flex-flex constraints arise from difficult cases of higher-order + unification. To prevent this, use @{ML res_inst_tac} to instantiate + some variables in a rule. Normally flex-flex constraints can be + ignored; they often disappear as unknowns get instantiated. + + \end{description} +*} + section {* Tacticals \label{sec:tacticals} *} +text {* A \emph{tactical} is a functional combinator for building up + complex tactics from simpler ones. Common tacticals perform + sequential composition, disjunctive choice, iteration, or goal + addressing. Various search strategies may be expressed via + tacticals. +*} + + +subsection {* Combining tactics *} + +text {* Sequential composition and alternative choices are the most + basic ways to combine tactics, similarly to ``@{verbatim ","}'' and + ``@{verbatim "|"}'' in Isar method notation. This corresponds to + @{ML_op "THEN"} and @{ML_op "ORELSE"} in ML, but there are further + possibilities for fine-tuning alternation of tactics such as @{ML_op + "APPEND"}. Further details become visible in ML due to explicit + subgoal addressing. +*} + +text %mlref {* + \begin{mldecls} + @{index_ML_op "THEN": "tactic * tactic -> tactic"} \\ + @{index_ML_op "ORELSE": "tactic * tactic -> tactic"} \\ + @{index_ML_op "APPEND": "tactic * tactic -> tactic"} \\ + @{index_ML "EVERY": "tactic list -> tactic"} \\ + @{index_ML "FIRST": "tactic list -> tactic"} \\[0.5ex] + + @{index_ML_op "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ + @{index_ML_op "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ + @{index_ML_op "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ + @{index_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\ + @{index_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{text "tac\<^sub>1"}~@{ML_op THEN}~@{text "tac\<^sub>2"} is the sequential + composition of @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Applied to a goal + state, it returns all states reachable in two steps by applying + @{text "tac\<^sub>1"} followed by @{text "tac\<^sub>2"}. First, it applies @{text + "tac\<^sub>1"} to the goal state, getting a sequence of possible next + states; then, it applies @{text "tac\<^sub>2"} to each of these and + concatenates the results to produce again one flat sequence of + states. + + \item @{text "tac\<^sub>1"}~@{ML_op ORELSE}~@{text "tac\<^sub>2"} makes a choice + between @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Applied to a state, it + tries @{text "tac\<^sub>1"} and returns the result if non-empty; if @{text + "tac\<^sub>1"} fails then it uses @{text "tac\<^sub>2"}. This is a deterministic + choice: if @{text "tac\<^sub>1"} succeeds then @{text "tac\<^sub>2"} is excluded + from the result. + + \item @{text "tac\<^sub>1"}~@{ML_op APPEND}~@{text "tac\<^sub>2"} concatenates the + possible results of @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"}. Unlike + @{ML_op "ORELSE"} there is \emph{no commitment} to either tactic, so + @{ML_op "APPEND"} helps to avoid incompleteness during search, at + the cost of potential inefficiencies. + + \item @{ML EVERY}~@{text "[tac\<^sub>1, \, tac\<^sub>n]"} abbreviates @{text + "tac\<^sub>1"}~@{ML_op THEN}~@{text "\"}~@{ML_op THEN}~@{text "tac\<^sub>n"}. + Note that @{ML "EVERY []"} is the same as @{ML all_tac}: it always + succeeds. + + \item @{ML FIRST}~@{text "[tac\<^sub>1, \, tac\<^sub>n]"} abbreviates @{text + "tac\<^sub>1"}~@{ML_op ORELSE}~@{text "\"}~@{ML_op "ORELSE"}~@{text + "tac\<^sub>n"}. Note that @{ML "FIRST []"} is the same as @{ML no_tac}: it + always fails. + + \item @{ML_op "THEN'"} is the lifted version of @{ML_op "THEN"}, for + tactics with explicit subgoal addressing. So @{text + "(tac\<^sub>1"}~@{ML_op THEN'}~@{text "tac\<^sub>2) i"} is the same as @{text + "(tac\<^sub>1 i"}~@{ML_op THEN}~@{text "tac\<^sub>2 i)"}. + + The other primed tacticals work analogously. + + \end{description} +*} + + +subsection {* Repetition tacticals *} + +text {* These tacticals provide further control over repetition of + tactics, beyond the stylized forms of ``@{verbatim "?"}'' and + ``@{verbatim "+"}'' in Isar method expressions. *} + +text %mlref {* + \begin{mldecls} + @{index_ML "TRY": "tactic -> tactic"} \\ + @{index_ML "REPEAT": "tactic -> tactic"} \\ + @{index_ML "REPEAT1": "tactic -> tactic"} \\ + @{index_ML "REPEAT_DETERM": "tactic -> tactic"} \\ + @{index_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML TRY}~@{text "tac"} applies @{text "tac"} to the goal + state and returns the resulting sequence, if non-empty; otherwise it + returns the original state. Thus, it applies @{text "tac"} at most + once. + + Note that for tactics with subgoal addressing, the combinator can be + applied via functional composition: @{ML "TRY"}~@{ML_op o}~@{text + "tac"}. There is no need for @{verbatim TRY'}. + + \item @{ML REPEAT}~@{text "tac"} applies @{text "tac"} to the goal + state and, recursively, to each element of the resulting sequence. + The resulting sequence consists of those states that make @{text + "tac"} fail. Thus, it applies @{text "tac"} as many times as + possible (including zero times), and allows backtracking over each + invocation of @{text "tac"}. @{ML REPEAT} is more general than @{ML + REPEAT_DETERM}, but requires more space. + + \item @{ML REPEAT1}~@{text "tac"} is like @{ML REPEAT}~@{text "tac"} + but it always applies @{text "tac"} at least once, failing if this + is impossible. + + \item @{ML REPEAT_DETERM}~@{text "tac"} applies @{text "tac"} to the + goal state and, recursively, to the head of the resulting sequence. + It returns the first state to make @{text "tac"} fail. It is + deterministic, discarding alternative outcomes. + + \item @{ML REPEAT_DETERM_N}~@{text "n tac"} is like @{ML + REPEAT_DETERM}~@{text "tac"} but the number of repetitions is bound + by @{text "n"} (where @{ML "~1"} means @{text "\"}). + + \end{description} +*} + +text %mlex {* The basic tactics and tacticals considered above follow + some algebraic laws: + + \begin{itemize} + + \item @{ML all_tac} is the identity element of the tactical @{ML_op + "THEN"}. + + \item @{ML no_tac} is the identity element of @{ML_op "ORELSE"} and + @{ML_op "APPEND"}. Also, it is a zero element for @{ML_op "THEN"}, + which means that @{text "tac"}~@{ML_op THEN}~@{ML no_tac} is + equivalent to @{ML no_tac}. + + \item @{ML TRY} and @{ML REPEAT} can be expressed as (recursive) + functions over more basic combinators (ignoring some internal + implementation tricks): + + \end{itemize} +*} + +ML {* + fun TRY tac = tac ORELSE all_tac; + fun REPEAT tac st = ((tac THEN REPEAT tac) ORELSE all_tac) st; +*} + +text {* If @{text "tac"} can return multiple outcomes then so can @{ML + REPEAT}~@{text "tac"}. @{ML REPEAT} uses @{ML_op "ORELSE"} and not + @{ML_op "APPEND"}, it applies @{text "tac"} as many times as + possible in each outcome. + + \begin{warn} + Note the explicit abstraction over the goal state in the ML + definition of @{ML REPEAT}. Recursive tacticals must be coded in + this awkward fashion to avoid infinite recursion of eager functional + evaluation in Standard ML. The following attempt would make @{ML + REPEAT}~@{text "tac"} loop: + \end{warn} +*} + +ML {* + (*BAD -- does not terminate!*) + fun REPEAT tac = (tac THEN REPEAT tac) ORELSE all_tac; +*} + + +subsection {* Applying tactics to subgoal ranges *} + +text {* Tactics with explicit subgoal addressing + @{ML_type "int -> tactic"} can be used together with tacticals that + act like ``subgoal quantifiers'': guided by success of the body + tactic a certain range of subgoals is covered. Thus the body tactic + is applied to \emph{all} subgoals, \emph{some} subgoal etc. + + Suppose that the goal state has @{text "n \ 0"} subgoals. Many of + these tacticals address subgoal ranges counting downwards from + @{text "n"} towards @{text "1"}. This has the fortunate effect that + newly emerging subgoals are concatenated in the result, without + interfering each other. Nonetheless, there might be situations + where a different order is desired. *} + +text %mlref {* + \begin{mldecls} + @{index_ML ALLGOALS: "(int -> tactic) -> tactic"} \\ + @{index_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\ + @{index_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\ + @{index_ML HEADGOAL: "(int -> tactic) -> tactic"} \\ + @{index_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\ + @{index_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\ + @{index_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML ALLGOALS}~@{text "tac"} is equivalent to @{text "tac + n"}~@{ML_op THEN}~@{text "\"}~@{ML_op THEN}~@{text "tac 1"}. It + applies the @{text tac} to all the subgoals, counting downwards. + + \item @{ML SOMEGOAL}~@{text "tac"} is equivalent to @{text "tac + n"}~@{ML_op ORELSE}~@{text "\"}~@{ML_op ORELSE}~@{text "tac 1"}. It + applies @{text "tac"} to one subgoal, counting downwards. + + \item @{ML FIRSTGOAL}~@{text "tac"} is equivalent to @{text "tac + 1"}~@{ML_op ORELSE}~@{text "\"}~@{ML_op ORELSE}~@{text "tac n"}. It + applies @{text "tac"} to one subgoal, counting upwards. + + \item @{ML HEADGOAL}~@{text "tac"} is equivalent to @{text "tac 1"}. + It applies @{text "tac"} unconditionally to the first subgoal. + + \item @{ML REPEAT_SOME}~@{text "tac"} applies @{text "tac"} once or + more to a subgoal, counting downwards. + + \item @{ML REPEAT_FIRST}~@{text "tac"} applies @{text "tac"} once or + more to a subgoal, counting upwards. + + \item @{ML RANGE}~@{text "[tac\<^sub>1, \, tac\<^sub>k] i"} is equivalent to + @{text "tac\<^sub>k (i + k - 1)"}~@{ML_op THEN}~@{text "\"}~@{ML_op + THEN}~@{text "tac\<^sub>1 i"}. It applies the given list of tactics to the + corresponding range of subgoals, counting downwards. + + \end{description} +*} + + +subsection {* Control and search tacticals *} + +text {* A predicate on theorems @{ML_type "thm -> bool"} can test + whether a goal state enjoys some desirable property --- such as + having no subgoals. Tactics that search for satisfactory goal + states are easy to express. The main search procedures, + depth-first, breadth-first and best-first, are provided as + tacticals. They generate the search tree by repeatedly applying a + given tactic. *} + + +text %mlref "" + +subsubsection {* Filtering a tactic's results *} + text {* - A \emph{tactical} is a functional combinator for building up complex - tactics from simpler ones. Typical tactical perform sequential - composition, disjunction (choice), iteration, or goal addressing. - Various search strategies may be expressed via tacticals. + \begin{mldecls} + @{index_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\ + @{index_ML CHANGED: "tactic -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML FILTER}~@{text "sat tac"} applies @{text "tac"} to the + goal state and returns a sequence consisting of those result goal + states that are satisfactory in the sense of @{text "sat"}. + + \item @{ML CHANGED}~@{text "tac"} applies @{text "tac"} to the goal + state and returns precisely those states that differ from the + original state (according to @{ML Thm.eq_thm}). Thus @{ML + CHANGED}~@{text "tac"} always has some effect on the state. + + \end{description} +*} + + +subsubsection {* Depth-first search *} + +text {* + \begin{mldecls} + @{index_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\ + @{index_ML DEPTH_SOLVE: "tactic -> tactic"} \\ + @{index_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML DEPTH_FIRST}~@{text "sat tac"} returns the goal state if + @{text "sat"} returns true. Otherwise it applies @{text "tac"}, + then recursively searches from each element of the resulting + sequence. The code uses a stack for efficiency, in effect applying + @{text "tac"}~@{ML_op THEN}~@{ML DEPTH_FIRST}~@{text "sat tac"} to + the state. + + \item @{ML DEPTH_SOLVE}@{text "tac"} uses @{ML DEPTH_FIRST} to + search for states having no subgoals. + + \item @{ML DEPTH_SOLVE_1}~@{text "tac"} uses @{ML DEPTH_FIRST} to + search for states having fewer subgoals than the given state. Thus, + it insists upon solving at least one subgoal. + + \end{description} +*} + + +subsubsection {* Other search strategies *} + +text {* + \begin{mldecls} + @{index_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\ + @{index_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\ + @{index_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\ + \end{mldecls} + + These search strategies will find a solution if one exists. + However, they do not enumerate all solutions; they terminate after + the first satisfactory result from @{text "tac"}. + + \begin{description} + + \item @{ML BREADTH_FIRST}~@{text "sat tac"} uses breadth-first + search to find states for which @{text "sat"} is true. For most + applications, it is too slow. + + \item @{ML BEST_FIRST}~@{text "(sat, dist) tac"} does a heuristic + search, using @{text "dist"} to estimate the distance from a + satisfactory state (in the sense of @{text "sat"}). It maintains a + list of states ordered by distance. It applies @{text "tac"} to the + head of this list; if the result contains any satisfactory states, + then it returns them. Otherwise, @{ML BEST_FIRST} adds the new + states to the list, and continues. - \medskip FIXME + The distance function is typically @{ML size_of_thm}, which computes + the size of the state. The smaller the state, the fewer and simpler + subgoals it has. + + \item @{ML THEN_BEST_FIRST}~@{text "tac\<^sub>0 (sat, dist) tac"} is like + @{ML BEST_FIRST}, except that the priority queue initially contains + the result of applying @{text "tac\<^sub>0"} to the goal state. This + tactical permits separate tactics for starting the search and + continuing the search. + + \end{description} +*} + + +subsubsection {* Auxiliary tacticals for searching *} + +text {* + \begin{mldecls} + @{index_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\ + @{index_ML IF_UNSOLVED: "tactic -> tactic"} \\ + @{index_ML SOLVE: "tactic -> tactic"} \\ + @{index_ML DETERM: "tactic -> tactic"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML COND}~@{text "sat tac\<^sub>1 tac\<^sub>2"} applies @{text "tac\<^sub>1"} to + the goal state if it satisfies predicate @{text "sat"}, and applies + @{text "tac\<^sub>2"}. It is a conditional tactical in that only one of + @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"} is applied to a goal state. + However, both @{text "tac\<^sub>1"} and @{text "tac\<^sub>2"} are evaluated + because ML uses eager evaluation. + + \item @{ML IF_UNSOLVED}~@{text "tac"} applies @{text "tac"} to the + goal state if it has any subgoals, and simply returns the goal state + otherwise. Many common tactics, such as @{ML resolve_tac}, fail if + applied to a goal state that has no subgoals. + + \item @{ML SOLVE}~@{text "tac"} applies @{text "tac"} to the goal + state and then fails iff there are subgoals left. - \medskip The chapter on tacticals in \cite{isabelle-ref} is still - applicable, despite a few outdated details. *} + \item @{ML DETERM}~@{text "tac"} applies @{text "tac"} to the goal + state and returns the head of the resulting sequence. @{ML DETERM} + limits the search space by making its argument deterministic. + + \end{description} +*} + + +subsubsection {* Predicates and functions useful for searching *} + +text {* + \begin{mldecls} + @{index_ML has_fewer_prems: "int -> thm -> bool"} \\ + @{index_ML Thm.eq_thm: "thm * thm -> bool"} \\ + @{index_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\ + @{index_ML size_of_thm: "thm -> int"} \\ + \end{mldecls} + + \begin{description} + + \item @{ML has_fewer_prems}~@{text "n thm"} reports whether @{text + "thm"} has fewer than @{text "n"} premises. + + \item @{ML Thm.eq_thm}~@{text "(thm\<^sub>1, thm\<^sub>2)"} reports whether @{text + "thm\<^sub>1"} and @{text "thm\<^sub>2"} are equal. Both theorems must have + compatible background theories. Both theorems must have the same + conclusions, the same set of hypotheses, and the same set of sort + hypotheses. Names of bound variables are ignored as usual. + + \item @{ML Thm.eq_thm_prop}~@{text "(thm\<^sub>1, thm\<^sub>2)"} reports whether + the propositions of @{text "thm\<^sub>1"} and @{text "thm\<^sub>2"} are equal. + Names of bound variables are ignored. + + \item @{ML size_of_thm}~@{text "thm"} computes the size of @{text + "thm"}, namely the number of variables, constants and abstractions + in its conclusion. It may serve as a distance function for + @{ML BEST_FIRST}. + + \end{description} +*} end diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/Eq.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/doc-src/IsarImplementation/Thy/document/Eq.tex Tue Feb 14 12:40:55 2012 +0100 @@ -0,0 +1,135 @@ +% +\begin{isabellebody}% +\def\isabellecontext{Eq}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{theory}\isamarkupfalse% +\ Eq\isanewline +\isakeyword{imports}\ Base\isanewline +\isakeyword{begin}% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +% +\isamarkupchapter{Equational reasoning% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Equality is one of the most fundamental concepts of + mathematics. The Isabelle/Pure logic (\chref{ch:logic}) provides a + builtin relation \isa{{\isaliteral{5C3C65717569763E}{\isasymequiv}}\ {\isaliteral{3A}{\isacharcolon}}{\isaliteral{3A}{\isacharcolon}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ {\isaliteral{5C3C616C7068613E}{\isasymalpha}}\ {\isaliteral{5C3C52696768746172726F773E}{\isasymRightarrow}}\ prop} that expresses equality + of arbitrary terms (or propositions) at the framework level, as + expressed by certain basic inference rules (\secref{sec:eq-rules}). + + Equational reasoning means to replace equals by equals, using + reflexivity and transitivity to form chains of replacement steps, + and congruence rules to access sub-structures. Conversions + (\secref{sec:conv}) provide a convenient framework to compose basic + equational steps to build specific equational reasoning tools. + + Higher-order matching is able to provide suitable instantiations for + giving equality rules, which leads to the versatile concept of + \isa{{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}}-term rewriting (\secref{sec:rewriting}). Internally + this is based on the general-purpose Simplifier engine of Isabelle, + which is more specific and more efficient than plain conversions. + + Object-logics usually introduce specific notions of equality or + equivalence, and relate it with the Pure equality. This enables to + re-use the Pure tools for equational reasoning for particular + object-logic connectives as well.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Basic equality rules \label{sec:eq-rules}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Conversions \label{sec:conv}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +FIXME% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Rewriting \label{sec:rewriting}% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Rewriting normalizes a given term (theorem or goal) by + replacing instances of given equalities \isa{t\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ u} in subterms. + Rewriting continues until no rewrites are applicable to any subterm. + This may be used to unfold simple definitions of the form \isa{f\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}\ x\isaliteral{5C3C5E7375623E}{}\isactrlsub n\ {\isaliteral{5C3C65717569763E}{\isasymequiv}}\ u}, but is slightly more general than that.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{rewrite\_goal\_tac}\verb|rewrite_goal_tac: thm list -> int -> tactic| \\ + \indexdef{}{ML}{rewrite\_goals\_tac}\verb|rewrite_goals_tac: thm list -> tactic| \\ + \indexdef{}{ML}{fold\_goals\_tac}\verb|fold_goals_tac: thm list -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|rewrite_goal_tac|~\isa{rules\ i} rewrites subgoal + \isa{i} by the given rewrite rules. + + \item \verb|rewrite_goals_tac|~\isa{rules} rewrites all subgoals + by the given rewrite rules. + + \item \verb|fold_goals_tac|~\isa{rules} essentially uses \verb|rewrite_goals_tac| with the symmetric form of each member of \isa{rules}, re-ordered to fold longer expression first. This supports + to idea to fold primitive definitions that appear in expended form + in the proof state. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimtheory +% +\endisadelimtheory +% +\isatagtheory +\isacommand{end}\isamarkupfalse% +% +\endisatagtheory +{\isafoldtheory}% +% +\isadelimtheory +% +\endisadelimtheory +\isanewline +\end{isabellebody}% +%%% Local Variables: +%%% mode: latex +%%% TeX-master: "root" +%%% End: diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/Isar.tex --- a/doc-src/IsarImplementation/Thy/document/Isar.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/document/Isar.tex Tue Feb 14 12:40:55 2012 +0100 @@ -428,7 +428,6 @@ \indexdef{}{ML}{METHOD}\verb|METHOD: (thm list -> tactic) -> Proof.method| \\ \indexdef{}{ML}{SIMPLE\_METHOD}\verb|SIMPLE_METHOD: tactic -> Proof.method| \\ \indexdef{}{ML}{SIMPLE\_METHOD'}\verb|SIMPLE_METHOD': (int -> tactic) -> Proof.method| \\ - \indexdef{}{ML}{HEADGOAL}\verb|HEADGOAL: (int -> tactic) -> tactic| \\ \indexdef{}{ML}{Method.insert\_tac}\verb|Method.insert_tac: thm list -> int -> tactic| \\ \indexdef{}{ML}{Method.setup}\verb|Method.setup: binding -> (Proof.context -> Proof.method) context_parser ->|\isasep\isanewline% \verb| string -> theory -> theory| \\ @@ -452,12 +451,8 @@ applied. \item \verb|SIMPLE_METHOD'|~\isa{tactic} wraps a tactic that - addresses a specific subgoal as simple proof method. Goal facts are - already inserted into the first subgoal before \isa{tactic} is - applied to the same. - - \item \verb|HEADGOAL|~\isa{tactic} applies \isa{tactic} to - the first subgoal. This is convenient to reproduce part the \verb|SIMPLE_METHOD'| wrapping within regular \verb|METHOD|, for example. + addresses a specific subgoal as simple proof method that operates on + subgoal 1. Goal facts are inserted into the subgoal then the \isa{tactic} is applied. \item \verb|Method.insert_tac|~\isa{facts\ i} inserts \isa{facts} into subgoal \isa{i}. This is convenient to reproduce part of the \verb|SIMPLE_METHOD| or \verb|SIMPLE_METHOD'| wrapping diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/Logic.tex --- a/doc-src/IsarImplementation/Thy/document/Logic.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/document/Logic.tex Tue Feb 14 12:40:55 2012 +0100 @@ -394,7 +394,7 @@ \begin{isamarkuptext}% \begin{mldecls} \indexdef{}{ML type}{term}\verb|type term| \\ - \indexdef{}{ML}{aconv}\verb|op aconv: term * term -> bool| \\ + \indexdef{}{ML infix}{aconv}\verb|infix aconv: term * term -> bool| \\ \indexdef{}{ML}{Term.map\_types}\verb|Term.map_types: (typ -> typ) -> term -> term| \\ \indexdef{}{ML}{Term.fold\_types}\verb|Term.fold_types: (typ -> 'a -> 'a) -> term -> 'a -> 'a| \\ \indexdef{}{ML}{Term.map\_aterms}\verb|Term.map_aterms: (term -> term) -> term -> term| \\ @@ -417,7 +417,7 @@ \item Type \verb|term| represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; - this is a datatype with constructors \verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|, \verb|Abs|, \verb|op $|. + this is a datatype with constructors \verb|Bound|, \verb|Free|, \verb|Var|, \verb|Const|, \verb|Abs|, \verb|$|. \item \isa{t}~\verb|aconv|~\isa{u} checks \isa{{\isaliteral{5C3C616C7068613E}{\isasymalpha}}}-equivalence of two terms. This is the basic equality relation on type \verb|term|; raw datatype equality should only be used @@ -704,10 +704,18 @@ % \begin{isamarkuptext}% \begin{mldecls} + \indexdef{}{ML}{Logic.all}\verb|Logic.all: term -> term -> term| \\ + \indexdef{}{ML}{Logic.mk\_implies}\verb|Logic.mk_implies: term * term -> term| \\ + \end{mldecls} + \begin{mldecls} \indexdef{}{ML type}{ctyp}\verb|type ctyp| \\ \indexdef{}{ML type}{cterm}\verb|type cterm| \\ \indexdef{}{ML}{Thm.ctyp\_of}\verb|Thm.ctyp_of: theory -> typ -> ctyp| \\ \indexdef{}{ML}{Thm.cterm\_of}\verb|Thm.cterm_of: theory -> term -> cterm| \\ + \indexdef{}{ML}{Thm.capply}\verb|Thm.capply: cterm -> cterm -> cterm| \\ + \indexdef{}{ML}{Thm.cabs}\verb|Thm.cabs: cterm -> cterm -> cterm| \\ + \indexdef{}{ML}{Thm.all}\verb|Thm.all: cterm -> cterm -> cterm| \\ + \indexdef{}{ML}{Drule.mk\_implies}\verb|Drule.mk_implies: cterm * cterm -> cterm| \\ \end{mldecls} \begin{mldecls} \indexdef{}{ML type}{thm}\verb|type thm| \\ @@ -734,19 +742,37 @@ \begin{description} + \item \verb|Logic.all|~\isa{a\ B} produces a Pure quantification + \isa{{\isaliteral{5C3C416E643E}{\isasymAnd}}a{\isaliteral{2E}{\isachardot}}\ B}, where occurrences of the atomic term \isa{a} in + the body proposition \isa{B} are replaced by bound variables. + (See also \verb|lambda| on terms.) + + \item \verb|Logic.mk_implies|~\isa{{\isaliteral{28}{\isacharparenleft}}A{\isaliteral{2C}{\isacharcomma}}\ B{\isaliteral{29}{\isacharparenright}}} produces a Pure + implication \isa{A\ {\isaliteral{5C3C4C6F6E6772696768746172726F773E}{\isasymLongrightarrow}}\ B}. + \item Types \verb|ctyp| and \verb|cterm| represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type - constructors, constants etc. in the theory. + constructors, constants etc.\ in the background theory. The + abstract types \verb|ctyp| and \verb|cterm| are part of the + same inference kernel that is mainly responsible for \verb|thm|. + Thus syntactic operations on \verb|ctyp| and \verb|cterm| + are located in the \verb|Thm| module, even though theorems are + not yet involved at that stage. \item \verb|Thm.ctyp_of|~\isa{thy\ {\isaliteral{5C3C7461753E}{\isasymtau}}} and \verb|Thm.cterm_of|~\isa{thy\ t} explicitly checks types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the theory context. + Full re-certification is relatively slow and should be avoided in + tight reasoning loops. - Re-certification is relatively slow and should be avoided in tight - reasoning loops. There are separate operations to decompose - certified entities (including actual theorems). + \item \verb|Thm.capply|, \verb|Thm.cabs|, \verb|Thm.all|, \verb|Drule.mk_implies| etc.\ compose certified terms (or propositions) + incrementally. This is equivalent to \verb|Thm.cterm_of| after + unchecked \verb|$|, \verb|lambda|, \verb|Logic.all|, \verb|Logic.mk_implies| etc., but there can be a big difference in + performance when large existing entities are composed by a few extra + constructions on top. There are separate operations to decompose + certified terms and theorems to produce certified terms again. \item Type \verb|thm| represents proven propositions. This is an abstract datatype that guarantees that its values have been @@ -917,14 +943,16 @@ % \endisadelimmlantiq % -\isamarkupsubsection{Auxiliary definitions \label{sec:logic-aux}% +\isamarkupsubsection{Auxiliary connectives \label{sec:logic-aux}% } \isamarkuptrue% % \begin{isamarkuptext}% -Theory \isa{Pure} provides a few auxiliary definitions, see - \figref{fig:pure-aux}. These special constants are normally not - exposed to the user, but appear in internal encodings. +Theory \isa{Pure} provides a few auxiliary connectives + that are defined on top of the primitive ones, see + \figref{fig:pure-aux}. These special constants are useful in + certain internal encodings, and are normally not directly exposed to + the user. \begin{figure}[htb] \begin{center} @@ -1198,21 +1226,46 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML}{RS}\verb|op RS: thm * thm -> thm| \\ - \indexdef{}{ML}{OF}\verb|op OF: thm * thm list -> thm| \\ + \indexdef{}{ML infix}{RSN}\verb|infix RSN: thm * (int * thm) -> thm| \\ + \indexdef{}{ML infix}{RS}\verb|infix RS: thm * thm -> thm| \\ + + \indexdef{}{ML infix}{RLN}\verb|infix RLN: thm list * (int * thm list) -> thm list| \\ + \indexdef{}{ML infix}{RL}\verb|infix RL: thm list * thm list -> thm list| \\ + + \indexdef{}{ML infix}{MRS}\verb|infix MRS: thm list * thm -> thm| \\ + \indexdef{}{ML infix}{OF}\verb|infix OF: thm * thm list -> thm| \\ \end{mldecls} \begin{description} - \item \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RS\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} resolves \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} with \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} according to the \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} principle - explained above. Note that the corresponding rule attribute in the - Isar language is called \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}}. + \item \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RSN\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{2C}{\isacharcomma}}\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}} resolves the conclusion of + \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} with the \isa{i}-th premise of \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}, + according to the \hyperlink{inference.resolution}{\mbox{\isa{resolution}}} principle explained above. + Unless there is precisely one resolvent it raises exception \verb|THM|. + + This corresponds to the rule attribute \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}} in Isar + source language. + + \item \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RS\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} abbreviates \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RS\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}}. - \item \isa{rule\ OF\ rules} resolves a list of rules with the - first rule, addressing its premises \isa{{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ length\ rules} - (operating from last to first). This means the newly emerging - premises are all concatenated, without interfering. Also note that - compared to \isa{RS}, the rule argument order is swapped: \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RS\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ {\isaliteral{3D}{\isacharequal}}\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ OF\ {\isaliteral{5B}{\isacharbrackleft}}rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{5D}{\isacharbrackright}}}. + \item \isa{rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RLN\ {\isaliteral{28}{\isacharparenleft}}i{\isaliteral{2C}{\isacharcomma}}\ rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}} joins lists of rules. For + every \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} in \isa{rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} in + \isa{rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}, it resolves the conclusion of \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} with + the \isa{i}-th premise of \isa{rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}, accumulating multiple + results in one big list. Note that such strict enumerations of + higher-order unifications can be inefficient compared to the lazy + variant seen in elementary tactics like \verb|resolve_tac|. + + \item \isa{rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RL\ rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} abbreviates \isa{rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ RLN\ {\isaliteral{28}{\isacharparenleft}}{\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ rules\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}}. + + \item \isa{{\isaliteral{5B}{\isacharbrackleft}}rule\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ rule\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}\ MRS\ rule} resolves \isa{rule\isaliteral{5C3C5E697375623E}{}\isactrlisub i} + against premise \isa{i} of \isa{rule}, for \isa{i\ {\isaliteral{3D}{\isacharequal}}\ n{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ {\isadigit{1}}}. By working from right to left, newly emerging premises are + concatenated in the result, without interfering. + + \item \isa{rule\ OF\ rules} abbreviates \isa{rules\ MRS\ rule}. + + This corresponds to the rule attribute \hyperlink{attribute.OF}{\mbox{\isa{OF}}} in Isar + source language. \end{description}% \end{isamarkuptext}% diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/ML.tex --- a/doc-src/IsarImplementation/Thy/document/ML.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/document/ML.tex Tue Feb 14 12:40:55 2012 +0100 @@ -1076,10 +1076,10 @@ % \begin{isamarkuptext}% \begin{mldecls} - \indexdef{}{ML}{$\mid$$>$}\verb|op |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\ - \indexdef{}{ML}{$\mid$-$>$}\verb|op |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\ - \indexdef{}{ML}{\#$>$}\verb|op #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\ - \indexdef{}{ML}{\#-$>$}\verb|op #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\ + \indexdef{}{ML infix}{$\mid$$>$}\verb|infix |\verb,|,\verb|> : 'a * ('a -> 'b) -> 'b| \\ + \indexdef{}{ML infix}{$\mid$-$>$}\verb|infix |\verb,|,\verb|-> : ('c * 'a) * ('c -> 'a -> 'b) -> 'b| \\ + \indexdef{}{ML infix}{\#$>$}\verb|infix #> : ('a -> 'b) * ('b -> 'c) -> 'a -> 'c| \\ + \indexdef{}{ML infix}{\#-$>$}\verb|infix #-> : ('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd| \\ \end{mldecls} %FIXME description!?% @@ -1968,7 +1968,7 @@ \indexdef{}{ML type}{Unsynchronized.ref}\verb|type 'a Unsynchronized.ref| \\ \indexdef{}{ML}{Unsynchronized.ref}\verb|Unsynchronized.ref: 'a -> 'a Unsynchronized.ref| \\ \indexdef{}{ML}{!}\verb|! : 'a Unsynchronized.ref -> 'a| \\ - \indexdef{}{ML}{:=}\verb|op := : 'a Unsynchronized.ref * 'a -> unit| \\ + \indexdef{}{ML infix}{:=}\verb|infix := : 'a Unsynchronized.ref * 'a -> unit| \\ \end{mldecls}% \end{isamarkuptext}% \isamarkuptrue% @@ -1989,7 +1989,7 @@ The unwieldy name of \verb|Unsynchronized.ref| for the constructor for references in Isabelle/ML emphasizes the inconveniences caused by - mutability. Existing operations \verb|!| and \verb|op :=| are + mutability. Existing operations \verb|!| and \verb|:=| are unchanged, but should be used with special precautions, say in a strictly local situation that is guaranteed to be restricted to sequential evaluation --- now and in the future. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/Proof.tex --- a/doc-src/IsarImplementation/Thy/document/Proof.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/document/Proof.tex Tue Feb 14 12:40:55 2012 +0100 @@ -588,6 +588,7 @@ % \begin{isamarkuptext}% \begin{mldecls} + \indexdef{}{ML}{SELECT\_GOAL}\verb|SELECT_GOAL: tactic -> int -> tactic| \\ \indexdef{}{ML}{SUBPROOF}\verb|SUBPROOF: (Subgoal.focus -> tactic) ->|\isasep\isanewline% \verb| Proof.context -> int -> tactic| \\ \indexdef{}{ML}{Subgoal.FOCUS}\verb|Subgoal.FOCUS: (Subgoal.focus -> tactic) ->|\isasep\isanewline% @@ -614,6 +615,10 @@ \begin{description} + \item \verb|SELECT_GOAL|~\isa{tac\ i} confines a tactic to the + specified subgoal \isa{i}. This introduces a nested goal state, + without decomposing the internal structure of the subgoal yet. + \item \verb|SUBPROOF|~\isa{tac\ ctxt\ i} decomposes the structure of the specified sub-goal, producing an extended context and a reduced goal, which needs to be solved by the given tactic. All diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/Thy/document/Tactic.tex --- a/doc-src/IsarImplementation/Thy/document/Tactic.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/Thy/document/Tactic.tex Tue Feb 14 12:40:55 2012 +0100 @@ -197,7 +197,7 @@ explicitly, and violating them merely results in ill-behaved tactics experienced by the user (e.g.\ tactics that insist in being applicable only to singleton goals, or prevent composition via - standard tacticals).% + standard tacticals such as \verb|REPEAT|).% \end{isamarkuptext}% \isamarkuptrue% % @@ -338,7 +338,10 @@ premises. \item \verb|eresolve_tac|~\isa{thms\ i} performs elim-resolution - with the given theorems, which should normally be elimination rules. + with the given theorems, which are normally be elimination rules. + + Note that \verb|eresolve_tac [asm_rl]| is equivalent to \verb|assume_tac|, which facilitates mixing of assumption steps with + genuine eliminations. \item \verb|dresolve_tac|~\isa{thms\ i} performs destruct-resolution with the given theorems, which should normally @@ -432,7 +435,9 @@ \indexdef{}{ML}{res\_inst\_tac}\verb|res_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ \indexdef{}{ML}{eres\_inst\_tac}\verb|eres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ \indexdef{}{ML}{dres\_inst\_tac}\verb|dres_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ - \indexdef{}{ML}{forw\_inst\_tac}\verb|forw_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\[1ex] + \indexdef{}{ML}{forw\_inst\_tac}\verb|forw_inst_tac: Proof.context -> (indexname * string) list -> thm -> int -> tactic| \\ + \indexdef{}{ML}{subgoal\_tac}\verb|subgoal_tac: Proof.context -> string -> int -> tactic| \\ + \indexdef{}{ML}{thin\_tac}\verb|thin_tac: Proof.context -> string -> int -> tactic| \\ \indexdef{}{ML}{rename\_tac}\verb|rename_tac: string list -> int -> tactic| \\ \end{mldecls} @@ -451,6 +456,17 @@ \item \verb|forw_inst_tac| is like \verb|dres_inst_tac| except that the selected assumption is not deleted. + \item \verb|subgoal_tac|~\isa{ctxt\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ i} adds the proposition + \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} as local premise to subgoal \isa{i}, and poses the + same as a new subgoal \isa{i\ {\isaliteral{2B}{\isacharplus}}\ {\isadigit{1}}} (in the original context). + + \item \verb|thin_tac|~\isa{ctxt\ {\isaliteral{5C3C7068693E}{\isasymphi}}\ i} deletes the specified + premise from subgoal \isa{i}. 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 \verb|rename_tac|~\isa{names\ i} renames the innermost parameters of subgoal \isa{i} according to the provided \isa{names} (which need to be distinct indentifiers). @@ -471,20 +487,571 @@ % \endisadelimmlref % +\isamarkupsubsection{Rearranging goal states% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +In rare situations there is a need to rearrange goal states: + either the overall collection of subgoals, or the local structure of + a subgoal. Various administrative tactics allow to operate on the + concrete presentation these conceptual sets of formulae.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{rotate\_tac}\verb|rotate_tac: int -> int -> tactic| \\ + \indexdef{}{ML}{distinct\_subgoals\_tac}\verb|distinct_subgoals_tac: tactic| \\ + \indexdef{}{ML}{flexflex\_tac}\verb|flexflex_tac: tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|rotate_tac|~\isa{n\ i} rotates the premises of subgoal + \isa{i} by \isa{n} positions: from right to left if \isa{n} is + positive, and from left to right if \isa{n} is negative. + + \item \verb|distinct_subgoals_tac| removes duplicate subgoals from a + proof state. This is potentially inefficient. + + \item \verb|flexflex_tac| removes all flex-flex pairs from the proof + state by applying the trivial unifier. This drastic step loses + information. It is already part of the Isar infrastructure for + facts resulting from goals, and rarely needs to be invoked manually. + + Flex-flex constraints arise from difficult cases of higher-order + unification. To prevent this, use \verb|res_inst_tac| to instantiate + some variables in a rule. Normally flex-flex constraints can be + ignored; they often disappear as unknowns get instantiated. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% \isamarkupsection{Tacticals \label{sec:tacticals}% } \isamarkuptrue% % \begin{isamarkuptext}% -A \emph{tactical} is a functional combinator for building up complex - tactics from simpler ones. Typical tactical perform sequential - composition, disjunction (choice), iteration, or goal addressing. - Various search strategies may be expressed via tacticals. +A \emph{tactical} is a functional combinator for building up + complex tactics from simpler ones. Common tacticals perform + sequential composition, disjunctive choice, iteration, or goal + addressing. Various search strategies may be expressed via + tacticals.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsection{Combining tactics% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Sequential composition and alternative choices are the most + basic ways to combine tactics, similarly to ``\verb|,|'' and + ``\verb||\verb,|,\verb||'' in Isar method notation. This corresponds to + \verb|THEN| and \verb|ORELSE| in ML, but there are further + possibilities for fine-tuning alternation of tactics such as \verb|APPEND|. Further details become visible in ML due to explicit + subgoal addressing.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML infix}{THEN}\verb|infix THEN: tactic * tactic -> tactic| \\ + \indexdef{}{ML infix}{ORELSE}\verb|infix ORELSE: tactic * tactic -> tactic| \\ + \indexdef{}{ML infix}{APPEND}\verb|infix APPEND: tactic * tactic -> tactic| \\ + \indexdef{}{ML}{EVERY}\verb|EVERY: tactic list -> tactic| \\ + \indexdef{}{ML}{FIRST}\verb|FIRST: tactic list -> tactic| \\[0.5ex] + + \indexdef{}{ML infix}{THEN'}\verb|infix THEN': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic| \\ + \indexdef{}{ML infix}{ORELSE'}\verb|infix ORELSE': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic| \\ + \indexdef{}{ML infix}{APPEND'}\verb|infix APPEND': ('a -> tactic) * ('a -> tactic) -> 'a -> tactic| \\ + \indexdef{}{ML}{EVERY'}\verb|EVERY': ('a -> tactic) list -> 'a -> tactic| \\ + \indexdef{}{ML}{FIRST'}\verb|FIRST': ('a -> tactic) list -> 'a -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|THEN|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} is the sequential + composition of \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. Applied to a goal + state, it returns all states reachable in two steps by applying + \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} followed by \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. First, it applies \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} to the goal state, getting a sequence of possible next + states; then, it applies \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} to each of these and + concatenates the results to produce again one flat sequence of + states. + + \item \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|ORELSE|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} makes a choice + between \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. Applied to a state, it + tries \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and returns the result if non-empty; if \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} fails then it uses \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. This is a deterministic + choice: if \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} succeeds then \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} is excluded + from the result. + + \item \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|APPEND|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} concatenates the + possible results of \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. Unlike + \verb|ORELSE| there is \emph{no commitment} to either tactic, so + \verb|APPEND| helps to avoid incompleteness during search, at + the cost of potential inefficiencies. + + \item \verb|EVERY|~\isa{{\isaliteral{5B}{\isacharbrackleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ tac\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}} abbreviates \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|THEN|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|THEN|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub n}. + Note that \verb|EVERY []| is the same as \verb|all_tac|: it always + succeeds. + + \item \verb|FIRST|~\isa{{\isaliteral{5B}{\isacharbrackleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ tac\isaliteral{5C3C5E7375623E}{}\isactrlsub n{\isaliteral{5D}{\isacharbrackright}}} abbreviates \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|ORELSE|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|ORELSE|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub n}. Note that \verb|FIRST []| is the same as \verb|no_tac|: it + always fails. + + \item \verb|THEN'| is the lifted version of \verb|THEN|, for + tactics with explicit subgoal addressing. So \isa{{\isaliteral{28}{\isacharparenleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}}~\verb|THEN'|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}\ i} is the same as \isa{{\isaliteral{28}{\isacharparenleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ i}~\verb|THEN|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}\ i{\isaliteral{29}{\isacharparenright}}}. + + The other primed tacticals work analogously. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Repetition tacticals% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +These tacticals provide further control over repetition of + tactics, beyond the stylized forms of ``\verb|?|'' and + ``\verb|+|'' in Isar method expressions.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{TRY}\verb|TRY: tactic -> tactic| \\ + \indexdef{}{ML}{REPEAT}\verb|REPEAT: tactic -> tactic| \\ + \indexdef{}{ML}{REPEAT1}\verb|REPEAT1: tactic -> tactic| \\ + \indexdef{}{ML}{REPEAT\_DETERM}\verb|REPEAT_DETERM: tactic -> tactic| \\ + \indexdef{}{ML}{REPEAT\_DETERM\_N}\verb|REPEAT_DETERM_N: int -> tactic -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|TRY|~\isa{tac} applies \isa{tac} to the goal + state and returns the resulting sequence, if non-empty; otherwise it + returns the original state. Thus, it applies \isa{tac} at most + once. + + Note that for tactics with subgoal addressing, the combinator can be + applied via functional composition: \verb|TRY|~\verb|o|~\isa{tac}. There is no need for \verb|TRY'|. + + \item \verb|REPEAT|~\isa{tac} applies \isa{tac} to the goal + state and, recursively, to each element of the resulting sequence. + The resulting sequence consists of those states that make \isa{tac} fail. Thus, it applies \isa{tac} as many times as + possible (including zero times), and allows backtracking over each + invocation of \isa{tac}. \verb|REPEAT| is more general than \verb|REPEAT_DETERM|, but requires more space. + + \item \verb|REPEAT1|~\isa{tac} is like \verb|REPEAT|~\isa{tac} + but it always applies \isa{tac} at least once, failing if this + is impossible. + + \item \verb|REPEAT_DETERM|~\isa{tac} applies \isa{tac} to the + goal state and, recursively, to the head of the resulting sequence. + It returns the first state to make \isa{tac} fail. It is + deterministic, discarding alternative outcomes. + + \item \verb|REPEAT_DETERM_N|~\isa{n\ tac} is like \verb|REPEAT_DETERM|~\isa{tac} but the number of repetitions is bound + by \isa{n} (where \verb|~1| means \isa{{\isaliteral{5C3C696E66696E6974793E}{\isasyminfinity}}}). + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isadelimmlex +% +\endisadelimmlex +% +\isatagmlex +% +\begin{isamarkuptext}% +The basic tactics and tacticals considered above follow + some algebraic laws: + + \begin{itemize} + + \item \verb|all_tac| is the identity element of the tactical \verb|THEN|. + + \item \verb|no_tac| is the identity element of \verb|ORELSE| and + \verb|APPEND|. Also, it is a zero element for \verb|THEN|, + which means that \isa{tac}~\verb|THEN|~\verb|no_tac| is + equivalent to \verb|no_tac|. + + \item \verb|TRY| and \verb|REPEAT| can be expressed as (recursive) + functions over more basic combinators (ignoring some internal + implementation tricks): + + \end{itemize}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlex +{\isafoldmlex}% +% +\isadelimmlex +% +\endisadelimmlex +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline +\ \ fun\ TRY\ tac\ {\isaliteral{3D}{\isacharequal}}\ tac\ ORELSE\ all{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{3B}{\isacharsemicolon}}\isanewline +\ \ fun\ REPEAT\ tac\ st\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{28}{\isacharparenleft}}tac\ THEN\ REPEAT\ tac{\isaliteral{29}{\isacharparenright}}\ ORELSE\ all{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{29}{\isacharparenright}}\ st{\isaliteral{3B}{\isacharsemicolon}}\isanewline +{\isaliteral{2A7D}{\isacharverbatimclose}}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\begin{isamarkuptext}% +If \isa{tac} can return multiple outcomes then so can \verb|REPEAT|~\isa{tac}. \verb|REPEAT| uses \verb|ORELSE| and not + \verb|APPEND|, it applies \isa{tac} as many times as + possible in each outcome. + + \begin{warn} + Note the explicit abstraction over the goal state in the ML + definition of \verb|REPEAT|. Recursive tacticals must be coded in + this awkward fashion to avoid infinite recursion of eager functional + evaluation in Standard ML. The following attempt would make \verb|REPEAT|~\isa{tac} loop: + \end{warn}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimML +% +\endisadelimML +% +\isatagML +\isacommand{ML}\isamarkupfalse% +\ {\isaliteral{7B2A}{\isacharverbatimopen}}\isanewline +\ \ {\isaliteral{28}{\isacharparenleft}}{\isaliteral{2A}{\isacharasterisk}}BAD\ {\isaliteral{2D}{\isacharminus}}{\isaliteral{2D}{\isacharminus}}\ does\ not\ terminate{\isaliteral{21}{\isacharbang}}{\isaliteral{2A}{\isacharasterisk}}{\isaliteral{29}{\isacharparenright}}\isanewline +\ \ fun\ REPEAT\ tac\ {\isaliteral{3D}{\isacharequal}}\ {\isaliteral{28}{\isacharparenleft}}tac\ THEN\ REPEAT\ tac{\isaliteral{29}{\isacharparenright}}\ ORELSE\ all{\isaliteral{5F}{\isacharunderscore}}tac{\isaliteral{3B}{\isacharsemicolon}}\isanewline +{\isaliteral{2A7D}{\isacharverbatimclose}}% +\endisatagML +{\isafoldML}% +% +\isadelimML +% +\endisadelimML +% +\isamarkupsubsection{Applying tactics to subgoal ranges% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +Tactics with explicit subgoal addressing + \verb|int -> tactic| can be used together with tacticals that + act like ``subgoal quantifiers'': guided by success of the body + tactic a certain range of subgoals is covered. Thus the body tactic + is applied to \emph{all} subgoals, \emph{some} subgoal etc. - \medskip FIXME + Suppose that the goal state has \isa{n\ {\isaliteral{5C3C67653E}{\isasymge}}\ {\isadigit{0}}} subgoals. Many of + these tacticals address subgoal ranges counting downwards from + \isa{n} towards \isa{{\isadigit{1}}}. This has the fortunate effect that + newly emerging subgoals are concatenated in the result, without + interfering each other. Nonetheless, there might be situations + where a different order is desired.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{ALLGOALS}\verb|ALLGOALS: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{SOMEGOAL}\verb|SOMEGOAL: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{FIRSTGOAL}\verb|FIRSTGOAL: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{HEADGOAL}\verb|HEADGOAL: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{REPEAT\_SOME}\verb|REPEAT_SOME: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{REPEAT\_FIRST}\verb|REPEAT_FIRST: (int -> tactic) -> tactic| \\ + \indexdef{}{ML}{RANGE}\verb|RANGE: (int -> tactic) list -> int -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|ALLGOALS|~\isa{tac} is equivalent to \isa{tac\ n}~\verb|THEN|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|THEN|~\isa{tac\ {\isadigit{1}}}. It + applies the \isa{tac} to all the subgoals, counting downwards. + + \item \verb|SOMEGOAL|~\isa{tac} is equivalent to \isa{tac\ n}~\verb|ORELSE|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|ORELSE|~\isa{tac\ {\isadigit{1}}}. It + applies \isa{tac} to one subgoal, counting downwards. + + \item \verb|FIRSTGOAL|~\isa{tac} is equivalent to \isa{tac\ {\isadigit{1}}}~\verb|ORELSE|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|ORELSE|~\isa{tac\ n}. It + applies \isa{tac} to one subgoal, counting upwards. + + \item \verb|HEADGOAL|~\isa{tac} is equivalent to \isa{tac\ {\isadigit{1}}}. + It applies \isa{tac} unconditionally to the first subgoal. + + \item \verb|REPEAT_SOME|~\isa{tac} applies \isa{tac} once or + more to a subgoal, counting downwards. + + \item \verb|REPEAT_FIRST|~\isa{tac} applies \isa{tac} once or + more to a subgoal, counting upwards. + + \item \verb|RANGE|~\isa{{\isaliteral{5B}{\isacharbrackleft}}tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ {\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{2C}{\isacharcomma}}\ tac\isaliteral{5C3C5E7375623E}{}\isactrlsub k{\isaliteral{5D}{\isacharbrackright}}\ i} is equivalent to + \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub k\ {\isaliteral{28}{\isacharparenleft}}i\ {\isaliteral{2B}{\isacharplus}}\ k\ {\isaliteral{2D}{\isacharminus}}\ {\isadigit{1}}{\isaliteral{29}{\isacharparenright}}}~\verb|THEN|~\isa{{\isaliteral{5C3C646F74733E}{\isasymdots}}}~\verb|THEN|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ i}. It applies the given list of tactics to the + corresponding range of subgoals, counting downwards. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsection{Control and search tacticals% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +A predicate on theorems \verb|thm -> bool| can test + whether a goal state enjoys some desirable property --- such as + having no subgoals. Tactics that search for satisfactory goal + states are easy to express. The main search procedures, + depth-first, breadth-first and best-first, are provided as + tacticals. They generate the search tree by repeatedly applying a + given tactic.% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isadelimmlref +% +\endisadelimmlref +% +\isatagmlref +% +\begin{isamarkuptext}% +% +\end{isamarkuptext}% +\isamarkuptrue% +% +\endisatagmlref +{\isafoldmlref}% +% +\isadelimmlref +% +\endisadelimmlref +% +\isamarkupsubsubsection{Filtering a tactic's results% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{FILTER}\verb|FILTER: (thm -> bool) -> tactic -> tactic| \\ + \indexdef{}{ML}{CHANGED}\verb|CHANGED: tactic -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|FILTER|~\isa{sat\ tac} applies \isa{tac} to the + goal state and returns a sequence consisting of those result goal + states that are satisfactory in the sense of \isa{sat}. + + \item \verb|CHANGED|~\isa{tac} applies \isa{tac} to the goal + state and returns precisely those states that differ from the + original state (according to \verb|Thm.eq_thm|). Thus \verb|CHANGED|~\isa{tac} always has some effect on the state. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Depth-first search% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{DEPTH\_FIRST}\verb|DEPTH_FIRST: (thm -> bool) -> tactic -> tactic| \\ + \indexdef{}{ML}{DEPTH\_SOLVE}\verb|DEPTH_SOLVE: tactic -> tactic| \\ + \indexdef{}{ML}{DEPTH\_SOLVE\_1}\verb|DEPTH_SOLVE_1: tactic -> tactic| \\ + \end{mldecls} - \medskip The chapter on tacticals in \cite{isabelle-ref} is still - applicable, despite a few outdated details.% + \begin{description} + + \item \verb|DEPTH_FIRST|~\isa{sat\ tac} returns the goal state if + \isa{sat} returns true. Otherwise it applies \isa{tac}, + then recursively searches from each element of the resulting + sequence. The code uses a stack for efficiency, in effect applying + \isa{tac}~\verb|THEN|~\verb|DEPTH_FIRST|~\isa{sat\ tac} to + the state. + + \item \verb|DEPTH_SOLVE|\isa{tac} uses \verb|DEPTH_FIRST| to + search for states having no subgoals. + + \item \verb|DEPTH_SOLVE_1|~\isa{tac} uses \verb|DEPTH_FIRST| to + search for states having fewer subgoals than the given state. Thus, + it insists upon solving at least one subgoal. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Other search strategies% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{BREADTH\_FIRST}\verb|BREADTH_FIRST: (thm -> bool) -> tactic -> tactic| \\ + \indexdef{}{ML}{BEST\_FIRST}\verb|BEST_FIRST: (thm -> bool) * (thm -> int) -> tactic -> tactic| \\ + \indexdef{}{ML}{THEN\_BEST\_FIRST}\verb|THEN_BEST_FIRST: tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic| \\ + \end{mldecls} + + These search strategies will find a solution if one exists. + However, they do not enumerate all solutions; they terminate after + the first satisfactory result from \isa{tac}. + + \begin{description} + + \item \verb|BREADTH_FIRST|~\isa{sat\ tac} uses breadth-first + search to find states for which \isa{sat} is true. For most + applications, it is too slow. + + \item \verb|BEST_FIRST|~\isa{{\isaliteral{28}{\isacharparenleft}}sat{\isaliteral{2C}{\isacharcomma}}\ dist{\isaliteral{29}{\isacharparenright}}\ tac} does a heuristic + search, using \isa{dist} to estimate the distance from a + satisfactory state (in the sense of \isa{sat}). It maintains a + list of states ordered by distance. It applies \isa{tac} to the + head of this list; if the result contains any satisfactory states, + then it returns them. Otherwise, \verb|BEST_FIRST| adds the new + states to the list, and continues. + + The distance function is typically \verb|size_of_thm|, which computes + the size of the state. The smaller the state, the fewer and simpler + subgoals it has. + + \item \verb|THEN_BEST_FIRST|~\isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}\ {\isaliteral{28}{\isacharparenleft}}sat{\isaliteral{2C}{\isacharcomma}}\ dist{\isaliteral{29}{\isacharparenright}}\ tac} is like + \verb|BEST_FIRST|, except that the priority queue initially contains + the result of applying \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{0}}} to the goal state. This + tactical permits separate tactics for starting the search and + continuing the search. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Auxiliary tacticals for searching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{COND}\verb|COND: (thm -> bool) -> tactic -> tactic -> tactic| \\ + \indexdef{}{ML}{IF\_UNSOLVED}\verb|IF_UNSOLVED: tactic -> tactic| \\ + \indexdef{}{ML}{SOLVE}\verb|SOLVE: tactic -> tactic| \\ + \indexdef{}{ML}{DETERM}\verb|DETERM: tactic -> tactic| \\ + \end{mldecls} + + \begin{description} + + \item \verb|COND|~\isa{sat\ tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}\ tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} applies \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} to + the goal state if it satisfies predicate \isa{sat}, and applies + \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}}. It is a conditional tactical in that only one of + \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} is applied to a goal state. + However, both \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{tac\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} are evaluated + because ML uses eager evaluation. + + \item \verb|IF_UNSOLVED|~\isa{tac} applies \isa{tac} to the + goal state if it has any subgoals, and simply returns the goal state + otherwise. Many common tactics, such as \verb|resolve_tac|, fail if + applied to a goal state that has no subgoals. + + \item \verb|SOLVE|~\isa{tac} applies \isa{tac} to the goal + state and then fails iff there are subgoals left. + + \item \verb|DETERM|~\isa{tac} applies \isa{tac} to the goal + state and returns the head of the resulting sequence. \verb|DETERM| + limits the search space by making its argument deterministic. + + \end{description}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsubsubsection{Predicates and functions useful for searching% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{mldecls} + \indexdef{}{ML}{has\_fewer\_prems}\verb|has_fewer_prems: int -> thm -> bool| \\ + \indexdef{}{ML}{Thm.eq\_thm}\verb|Thm.eq_thm: thm * thm -> bool| \\ + \indexdef{}{ML}{Thm.eq\_thm\_prop}\verb|Thm.eq_thm_prop: thm * thm -> bool| \\ + \indexdef{}{ML}{size\_of\_thm}\verb|size_of_thm: thm -> int| \\ + \end{mldecls} + + \begin{description} + + \item \verb|has_fewer_prems|~\isa{n\ thm} reports whether \isa{thm} has fewer than \isa{n} premises. + + \item \verb|Thm.eq_thm|~\isa{{\isaliteral{28}{\isacharparenleft}}thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}} reports whether \isa{thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} are equal. Both theorems must have + compatible background theories. Both theorems must have the same + conclusions, the same set of hypotheses, and the same set of sort + hypotheses. Names of bound variables are ignored as usual. + + \item \verb|Thm.eq_thm_prop|~\isa{{\isaliteral{28}{\isacharparenleft}}thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}{\isaliteral{2C}{\isacharcomma}}\ thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}{\isaliteral{29}{\isacharparenright}}} reports whether + the propositions of \isa{thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{1}}} and \isa{thm\isaliteral{5C3C5E7375623E}{}\isactrlsub {\isadigit{2}}} are equal. + Names of bound variables are ignored. + + \item \verb|size_of_thm|~\isa{thm} computes the size of \isa{thm}, namely the number of variables, constants and abstractions + in its conclusion. It may serve as a distance function for + \verb|BEST_FIRST|. + + \end{description}% \end{isamarkuptext}% \isamarkuptrue% % diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarImplementation/implementation.tex --- a/doc-src/IsarImplementation/implementation.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarImplementation/implementation.tex Tue Feb 14 12:40:55 2012 +0100 @@ -86,6 +86,7 @@ \input{Thy/document/Logic.tex} \input{Thy/document/Syntax.tex} \input{Thy/document/Tactic.tex} +\input{Thy/document/Eq.tex} \input{Thy/document/Proof.tex} \input{Thy/document/Isar.tex} \input{Thy/document/Local_Theory.tex} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Document_Preparation.thy --- a/doc-src/IsarRef/Thy/Document_Preparation.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Document_Preparation.thy Tue Feb 14 12:40:55 2012 +0100 @@ -155,6 +155,7 @@ @{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} \\ @@ -197,12 +198,15 @@ @@{antiquotation typ} options @{syntax type} | @@{antiquotation type} options @{syntax name} | @@{antiquotation class} options @{syntax name} | - @@{antiquotation text} 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} @@ -289,9 +293,10 @@ information omitted in the compact proof term, which is denoted by ``@{text _}'' placeholders there. - \item @{text "@{ML s}"}, @{text "@{ML_type s}"}, and @{text - "@{ML_struct s}"} check text @{text s} as ML value, type, and - structure, respectively. The source is printed verbatim. + \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. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Generic.thy --- a/doc-src/IsarRef/Thy/Generic.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Generic.thy Tue Feb 14 12:40:55 2012 +0100 @@ -160,8 +160,8 @@ compose rules by resolution. @{attribute THEN} resolves with the first premise of @{text a} (an alternative position may be also specified); the @{attribute COMP} version skips the automatic - lifting process that is normally intended (cf.\ @{ML "op RS"} and - @{ML "op COMP"} in \cite{isabelle-implementation}). + lifting process that is normally intended (cf.\ @{ML_op "RS"} and + @{ML_op "COMP"} 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 @@ -335,23 +335,25 @@ may be given as well, see also ML tactic @{ML cut_inst_tac} in \cite{isabelle-implementation}. - \item @{method thin_tac}~@{text \} deletes the specified assumption - from a subgoal; note that @{text \} may contain schematic variables. - See also @{ML thin_tac} in \cite{isabelle-implementation}. + \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 \} adds @{text \} as an - assumption to a subgoal. See also @{ML subgoal_tac} and @{ML - subgoals_tac} in \cite{isabelle-implementation}. + \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 assumptions of a - goal by @{text n} positions: from right to left if @{text n} is + \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. See also @{ML rotate_tac} in - \cite{isabelle-implementation}. + 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 diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/HOL_Specific.thy --- a/doc-src/IsarRef/Thy/HOL_Specific.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/HOL_Specific.thy Tue Feb 14 12:40:55 2012 +0100 @@ -60,11 +60,20 @@ section {* Inductive and coinductive definitions \label{sec:hol-inductive} *} -text {* 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. +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: @@ -86,14 +95,6 @@ introduction rule. The default rule declarations of Isabelle/HOL already take care of most common situations. - \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} - @{rail " (@@{command (HOL) inductive} | @@{command (HOL) inductive_set} | @@{command (HOL) coinductive} | @@{command (HOL) coinductive_set}) @@ -545,10 +546,10 @@ 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. + 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} *} @@ -594,12 +595,13 @@ 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"}. + 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 @@ -609,6 +611,7 @@ 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 @@ -626,15 +629,15 @@ subsection {* Old-style recursive function definitions (TFL) *} text {* - 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. - \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? @@ -1057,8 +1060,13 @@ section {* Typedef axiomatization \label{sec:hol-typedef} *} -text {* A Gordon/HOL-style type definition is a certain axiom scheme - that identifies a new type with a subset of an existing type. More +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 @@ -1078,10 +1086,6 @@ type_synonym} from Isabelle/Pure merely introduces syntactic abbreviations, without any logical significance. - \begin{matharray}{rcl} - @{command_def (HOL) "typedef"} & : & @{text "local_theory \ proof(prove)"} \\ - \end{matharray} - @{rail " @@{command (HOL) typedef} alt_name? abs_type '=' rep_set ; @@ -1251,12 +1255,6 @@ section {* Quotient types *} text {* - 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{matharray}{rcl} @{command_def (HOL) "quotient_type"} & : & @{text "local_theory \ proof(prove)"}\\ @{command_def (HOL) "quotient_definition"} & : & @{text "local_theory \ proof(prove)"}\\ @@ -1278,6 +1276,12 @@ @{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'); @@ -1303,17 +1307,22 @@ \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 + \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_quotmaps"} prints quotient map functions. + \item @{command (HOL) "quotient_definition"} defines a constant on + the quotient type. + + \item @{command (HOL) "print_quotmaps"} prints quotient map + functions. \item @{command (HOL) "print_quotients"} prints quotients. @@ -1379,9 +1388,9 @@ theorems. \end{description} - *} + section {* Coercive subtyping *} text {* @@ -1391,6 +1400,11 @@ @{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})? ; @@ -1400,45 +1414,36 @@ ; "} - 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{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 nullary type constructors. 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 + 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. - + 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 {* @@ -1448,18 +1453,22 @@ @{attribute_def (HOL) arith_split} & : & @{text attribute} \\ \end{matharray} - The @{method (HOL) arith} method decides linear arithmetic problems - (on types @{text nat}, @{text int}, @{text real}). Any current - facts are inserted into the goal before running the procedure. - - The @{attribute (HOL) arith} attribute declares facts that are - always supplied to the arithmetic provers implicitly. - - The @{attribute (HOL) arith_split} attribute declares case split + \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. - Note that a simpler (but faster) arithmetic prover is - already invoked by the Simplifier. + \end{description} + + Note that a simpler (but faster) arithmetic prover is already + invoked by the Simplifier. *} @@ -1474,10 +1483,12 @@ @@{method (HOL) iprover} ( @{syntax rulemod} * ) "} - The @{method (HOL) iprover} method 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. + \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 @@ -1487,8 +1498,11 @@ 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 {* @@ -1501,22 +1515,28 @@ @@{method (HOL) meson} @{syntax thmrefs}? ; - @@{method (HOL) metis} ( '(' ('partial_types' | 'full_types' | 'no_types' - | @{syntax name}) ')' )? @{syntax thmrefs}? + @@{method (HOL) metis} + ('(' ('partial_types' | 'full_types' | 'no_types' | @{syntax name}) ')')? + @{syntax thmrefs}? "} - The @{method (HOL) meson} method implements Loveland's model elimination - procedure \cite{loveland-78}. See @{file "~~/src/HOL/ex/Meson_Test.thy"} for - examples. - - The @{method (HOL) metis} method 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 @{file - "~~/src/HOL/Metis_Examples"} directory contains several small theories - developed to a large extent using Metis. + \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 {* @@ -1528,11 +1548,14 @@ @@{method (HOL) coherent} @{syntax thmrefs}? "} - The @{method (HOL) coherent} method 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. + \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} *} @@ -1574,27 +1597,29 @@ \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) "try_methods"} attempts to prove a subgoal using a combination - of standard proof methods (@{text auto}, @{text simp}, @{text 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) "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) "try_methods"} 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 (@{text "solve_direct"}, - @{text "quickcheck"}, @{text "try_methods"}, @{text "sledgehammer"}, - @{text "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. + using a combination of provers and disprovers (@{command (HOL) + "solve_direct"}, @{command (HOL) "quickcheck"}, @{command (HOL) + "try_methods"}, @{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} *} @@ -1647,54 +1672,54 @@ \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}. + 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: + 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 - @{text quickcheck_exhaustive_active}, @{text quickcheck_random_active}, - @{text quickcheck_narrowing_active} are set to true. + 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. + counterexample, but not potentially spurious counterexamples due + to underspecified functions. \item[@{text eval}] takes a term or a list of terms and evaluates - these terms under the variable assignment found by quickcheck. + these terms under the variable assignment found by quickcheck. \item[@{text iterations}] sets how many sets of assignments are generated for each particular size. @@ -1713,8 +1738,8 @@ \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 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 @@ -1722,31 +1747,31 @@ \end{description} - These option can be given within square brackets. - - \item @{command (HOL) "quickcheck_params"} changes - @{command (HOL) "quickcheck"} configuration options persistently. + 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. + 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: + 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 $\infty$. - - \item[@{text maxvars}] specifies the maximum number of Boolean variables - to use when transforming the term into a propositional formula. - Nonpositive values mean $\infty$. + \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. @@ -1756,22 +1781,22 @@ \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}). + 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. + 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. \end{description} *} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Inner_Syntax.thy --- a/doc-src/IsarRef/Thy/Inner_Syntax.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Inner_Syntax.thy Tue Feb 14 12:40:55 2012 +0100 @@ -4,9 +4,35 @@ 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 + translations} --- either as rewrite systems on first-order ASTs + (\secref{sec:syn-trans}) or ML functions on ASTs or @{text + "\"}-terms that represent parse trees (\secref{sec:tr-funs}). +*} + + section {* Printing logical entities *} -subsection {* Diagnostic commands *} +subsection {* Diagnostic commands \label{sec:print-diag} *} text {* \begin{matharray}{rcl} @@ -78,11 +104,11 @@ All of the diagnostic commands above admit a list of @{text modes} to be specified, which is appended to the current print mode; see - \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 + 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 @@ -206,6 +232,77 @@ *} +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 {* @@ -234,49 +331,63 @@ *} -section {* Mixfix annotations *} +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 specifications etc.\ also admit mixfix - annotations. + 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 \"infix\"}: '(' (@'infix' | @'infixl' | @'infixr') - @{syntax string} @{syntax nat} ')' + @{syntax_def mixfix}: '(' mfix ')' ; - @{syntax_def mixfix}: @{syntax \"infix\"} | '(' @{syntax string} prios? @{syntax nat}? ')' | - '(' @'binder' @{syntax string} prios? @{syntax nat} ')' - ; - @{syntax_def struct_mixfix}: @{syntax mixfix} | '(' @'structure' ')' + @{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} + ',') ']' "} - Here the @{syntax string} specifications refer to the actual mixfix - template, which 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. + 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. +*} - \medskip 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. + +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. + 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 @@ -340,17 +451,85 @@ \end{description} - For example, the template @{verbatim "(_ +/ _)"} specifies an infix - operator. There are two argument positions; the delimiter - @{verbatim "+"} is preceded by a space and followed by a space or - line break; the entire phrase is a pretty printing block. + 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} - The general idea of pretty printing with blocks and breaks is also - described in \cite{paulson-ml2}. + @{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. *} -section {* Explicit notation *} +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} @@ -361,6 +540,13 @@ @{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') @@ -376,7 +562,7 @@ \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. @@ -384,7 +570,7 @@ \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. @@ -393,15 +579,64 @@ works within an Isar proof body. \end{description} - - Note that the more primitive commands @{command "syntax"} and - @{command "no_syntax"} (\secref{sec:syn-trans}) provide raw access - to the syntax tables of a global theory. *} 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) xstr} & = & @{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) + xstr} 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 @@ -477,13 +712,10 @@ *} -subsection {* The Pure grammar *} +subsection {* The Pure grammar \label{sec:pure-grammar} *} -text {* - The priority grammar of the @{text "Pure"} theory is defined as follows: - - %FIXME syntax for "index" (?) - %FIXME "op" versions of ==> etc. (?) +text {* The priority grammar of the @{text "Pure"} theory is defined + approximately like this: \begin{center} \begin{supertabular}{rclr} @@ -509,21 +741,28 @@ @{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} \\\\ @@ -533,16 +772,17 @@ @{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 "id | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) id | "}@{verbatim "("} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim ")"} @{text id} \\ - & @{text "|"} & @{text "longid | type\<^sup>(\<^sup>1\<^sup>0\<^sup>0\<^sup>0\<^sup>) longid"} \\ - & @{text "|"} & @{verbatim "("} @{text type} @{verbatim ","} @{text "\"} @{verbatim ","} @{text type} @{verbatim ")"} @{text longid} \\ + & @{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)"} \\\\ + & @{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} & = & @{text "id | longid | "}@{verbatim "{}"} \\ - & @{text "|"} & @{verbatim "{"} @{text "(id | longid)"} @{verbatim ","} @{text "\"} @{verbatim ","} @{text "(id | longid)"} @{verbatim "}"} \\ + @{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} @@ -581,6 +821,12 @@ (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. @@ -657,178 +903,20 @@ 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} *} -section {* 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) xstr} & = & @{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) - xstr} 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. -*} - - -section {* Syntax and translations \label{sec:syn-trans} *} - -text {* - \begin{matharray}{rcl} - @{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"} \\ - \end{matharray} - - @{rail " - @@{command nonterminal} (@{syntax name} + @'and') - ; - (@@{command syntax} | @@{command no_syntax}) @{syntax mode}? (@{syntax constdecl} +) - ; - (@@{command translations} | @@{command no_translations}) - (transpat ('==' | '=>' | '<=' | '\' | '\' | '\') transpat +) - ; - - 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) decls"} is similar to - @{command "consts"}~@{text decls}, except that the actual logical - signature extension is omitted. Thus the context free grammar of - Isabelle's inner syntax may be augmented in arbitrary ways, - independently of the logic. The @{text mode} argument refers to the - print mode that the grammar rules belong; unless the @{keyword_ref - "output"} indicator is given, all productions are added both to the - input and output grammar. - - \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): parse~/ print rules (@{text "\"}), - parse rules (@{text "\"}), or print rules (@{text "\"}). - Translation patterns may be prefixed by the syntactic category to be - used for parsing; the default is @{text logic}. - - \item @{command "no_translations"}~@{text rules} removes syntactic - translation rules, which are interpreted in the same manner as for - @{command "translations"} above. - - \end{description} -*} - - -section {* 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"} \\ - \end{matharray} - - @{rail " - ( @@{command parse_ast_translation} | @@{command parse_translation} | - @@{command print_translation} | @@{command typed_print_translation} | - @@{command print_ast_translation}) ('(' @'advanced' ')')? @{syntax text} - "} - - Syntax translation functions written in ML admit almost arbitrary - manipulations of Isabelle's inner syntax. Any of the above commands - have a single @{syntax text} argument that refers to an ML - expression of appropriate type, which are as follows by default: - -%FIXME proper antiquotations -\begin{ttbox} -val parse_ast_translation : (string * (ast list -> ast)) list -val parse_translation : (string * (term list -> term)) list -val print_translation : (string * (term list -> term)) list -val typed_print_translation : (string * (typ -> term list -> term)) list -val print_ast_translation : (string * (ast list -> ast)) list -\end{ttbox} - - If the @{text "(advanced)"} option is given, the corresponding - translation functions may depend on the current theory or proof - context. This allows to implement advanced syntax mechanisms, as - translations functions may refer to specific theory declarations or - auxiliary proof data. - - See also \cite{isabelle-ref} for more information on the general - concept of syntax transformations in Isabelle. - -%FIXME proper antiquotations -\begin{ttbox} -val parse_ast_translation: - (string * (Proof.context -> ast list -> ast)) list -val parse_translation: - (string * (Proof.context -> term list -> term)) list -val print_translation: - (string * (Proof.context -> term list -> term)) list -val typed_print_translation: - (string * (Proof.context -> typ -> term list -> term)) list -val print_ast_translation: - (string * (Proof.context -> ast list -> ast)) list -\end{ttbox} -*} - - -section {* Inspecting the syntax *} +subsection {* Inspecting the syntax *} text {* \begin{matharray}{rcl} @@ -883,8 +971,231 @@ functions; see \secref{sec:tr-funs}. \end{description} - + + \end{description} +*} + + +subsection {* Ambiguity of parsed expressions *} + +text {* + \begin{tabular}{rcll} + @{attribute_def syntax_ambiguity_level} & : & @{text attribute} & default @{text 1} \\ + \end{tabular} + + \begin{mldecls} + @{index_ML Syntax.ambiguity_limit: "int Config.T"} \\ %FIXME attribute + \end{mldecls} + + 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_level} determines the number of + parser results that are tolerated without printing a detailed + message. + + \item @{ML 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 {* Raw syntax and translations \label{sec:syn-trans} *} + +text {* + \begin{matharray}{rcl} + @{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"} \\ + \end{matharray} + + Unlike mixfix notation for existing formal entities + (\secref{sec:notation}), raw syntax declarations provide full access + to the priority grammar of the inner syntax. 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}? (@{syntax constdecl} +) + ; + (@@{command translations} | @@{command no_translations}) + (transpat ('==' | '=>' | '<=' | '\' | '\' | '\') transpat +) + ; + + 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 unrelated 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): parse~/ print rules (@{text "\"}), + parse rules (@{text "\"}), or print rules (@{text "\"}). + Translation patterns may be prefixed by the syntactic category to be + used for parsing; the default is @{text logic}. + + \item @{command "no_translations"}~@{text rules} removes syntactic + translation rules, which are interpreted in the same manner as for + @{command "translations"} above. + + \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} +*} + + +section {* 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"} \\ + \end{matharray} + + @{rail " + ( @@{command parse_ast_translation} | @@{command parse_translation} | + @@{command print_translation} | @@{command typed_print_translation} | + @@{command print_ast_translation}) ('(' @'advanced' ')')? @{syntax text} + "} + + Syntax translation functions written in ML admit almost arbitrary + manipulations of Isabelle's inner syntax. Any of the above commands + have a single @{syntax text} argument that refers to an ML + expression of appropriate type, which are as follows by default: + +%FIXME proper antiquotations +\begin{ttbox} +val parse_ast_translation : (string * (ast list -> ast)) list +val parse_translation : (string * (term list -> term)) list +val print_translation : (string * (term list -> term)) list +val typed_print_translation : (string * (typ -> term list -> term)) list +val print_ast_translation : (string * (ast list -> ast)) list +\end{ttbox} + + If the @{text "(advanced)"} option is given, the corresponding + translation functions may depend on the current theory or proof + context. This allows to implement advanced syntax mechanisms, as + translations functions may refer to specific theory declarations or + auxiliary proof data. + +%FIXME proper antiquotations +\begin{ttbox} +val parse_ast_translation: + (string * (Proof.context -> ast list -> ast)) list +val parse_translation: + (string * (Proof.context -> term list -> term)) list +val print_translation: + (string * (Proof.context -> term list -> term)) list +val typed_print_translation: + (string * (Proof.context -> typ -> term list -> term)) list +val print_ast_translation: + (string * (Proof.context -> ast list -> ast)) list +\end{ttbox} + + \medskip See also the chapter on ``Syntax Transformations'' in old + \cite{isabelle-ref} for further details on translations on parse + trees. +*} + end diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/ML_Tactic.thy --- a/doc-src/IsarRef/Thy/ML_Tactic.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/ML_Tactic.thy Tue Feb 14 12:40:55 2012 +0100 @@ -103,14 +103,11 @@ 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.\ (see \cite{isabelle-ref}). The corresponding - Isar methods usually share the same base name, such as @{method - blast}, @{method fast}, @{method clarify} etc.\ (see - \secref{sec:classical}). -*} +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 *} @@ -159,25 +156,25 @@ \end{tabular} \medskip - \medskip @{ML CHANGED} (see \cite{isabelle-ref}) 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 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-ref}) 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 + \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, such as solving - goals in a different order. + 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}). + \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 ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Misc.thy --- a/doc-src/IsarRef/Thy/Misc.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Misc.thy Tue Feb 14 12:40:55 2012 +0100 @@ -114,35 +114,6 @@ *} -section {* History commands \label{sec:history} *} - -text {* - \begin{matharray}{rcl} - @{command_def "undo"}^{{ * }{ * }} & : & @{text "any \ any"} \\ - @{command_def "linear_undo"}^{{ * }{ * }} & : & @{text "any \ any"} \\ - @{command_def "kill"}^{{ * }{ * }} & : & @{text "any \ any"} \\ - \end{matharray} - - The Isabelle/Isar top-level maintains a two-stage history, for - theory and proof state transformation. Basically, any command can - be undone using @{command "undo"}, excluding mere diagnostic - elements. Note that a theorem statement with a \emph{finished} - proof is treated as a single unit by @{command "undo"}. In - contrast, the variant @{command "linear_undo"} admits to step back - into the middle of a proof. The @{command "kill"} command aborts - the current history node altogether, discontinuing a proof or even - the whole theory. This operation is \emph{not} undo-able. - - \begin{warn} - History commands should never be used with user interfaces such as - Proof~General \cite{proofgeneral,Aspinall:TACAS:2000}, which takes - care of stepping forth and back itself. Interfering by manual - @{command "undo"}, @{command "linear_undo"}, or even @{command - "kill"} commands would quickly result in utter confusion. - \end{warn} -*} - - section {* System commands *} text {* diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Outer_Syntax.thy --- a/doc-src/IsarRef/Thy/Outer_Syntax.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Outer_Syntax.thy Tue Feb 14 12:40:55 2012 +0100 @@ -2,7 +2,7 @@ imports Base Main begin -chapter {* Outer syntax *} +chapter {* Outer syntax --- the theory language *} text {* The rather generic framework of Isabelle/Isar syntax emerges from diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/Proof.thy --- a/doc-src/IsarRef/Thy/Proof.thy Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/Proof.thy Tue Feb 14 12:40:55 2012 +0100 @@ -138,10 +138,6 @@ logically sound manner. Note that the Isabelle {\LaTeX} macros can be easily adapted to print something like ``@{text "\"}'' instead of the keyword ``@{command "oops"}''. - - \medskip The @{command "oops"} command is undo-able, unlike - @{command_ref "kill"} (see \secref{sec:history}). The effect is to - get back to the theory just before the opening of the proof. *} @@ -790,7 +786,7 @@ \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 parallel). This corresponds to the @{ML "op MRS"} operation in + (in parallel). This corresponds to the @{ML_op "MRS"} operation in ML, but note the reversed order. Positions may be effectively skipped by including ``@{text _}'' (underscore) as argument. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Document_Preparation.tex --- a/doc-src/IsarRef/Thy/document/Document_Preparation.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Document_Preparation.tex Tue Feb 14 12:40:55 2012 +0100 @@ -205,6 +205,7 @@ \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} \\ @@ -236,7 +237,7 @@ \rail@nont{\isa{antiquotation}}[] \rail@term{\isa{{\isaliteral{7D}{\isacharbraceright}}}}[] \rail@end -\rail@begin{23}{\indexdef{}{syntax}{antiquotation}\hypertarget{syntax.antiquotation}{\hyperlink{syntax.antiquotation}{\mbox{\isa{antiquotation}}}}} +\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}}[] @@ -305,33 +306,40 @@ \rail@term{\hyperlink{antiquotation.text}{\mbox{\isa{text}}}}[] \rail@nont{\isa{options}}[] \rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] -\rail@nextbar{15} +\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{16} +\rail@nextbar{1} \rail@term{\hyperlink{antiquotation.subgoals}{\mbox{\isa{subgoals}}}}[] \rail@nont{\isa{options}}[] -\rail@nextbar{17} +\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{18} +\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{19} +\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{20} +\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{21} +\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{22} +\rail@nextbar{8} \rail@term{\hyperlink{antiquotation.file}{\mbox{\isa{file}}}}[] \rail@nont{\isa{options}}[] \rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[] @@ -445,8 +453,9 @@ 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}}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, type, and - structure, respectively. The source is printed verbatim. + \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. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Generic.tex --- a/doc-src/IsarRef/Thy/document/Generic.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Generic.tex Tue Feb 14 12:40:55 2012 +0100 @@ -277,8 +277,8 @@ compose rules by resolution. \hyperlink{attribute.THEN}{\mbox{\isa{THEN}}} resolves with the first premise of \isa{a} (an alternative position may be also specified); the \hyperlink{attribute.COMP}{\mbox{\isa{COMP}}} version skips the automatic - lifting process that is normally intended (cf.\ \verb|op RS| and - \verb|op COMP| in \cite{isabelle-implementation}). + lifting process that is normally intended (cf.\ \verb|RS| and + \verb|COMP| 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. @@ -537,22 +537,25 @@ may be given as well, see also ML tactic \verb|cut_inst_tac| in \cite{isabelle-implementation}. - \item \hyperlink{method.thin-tac}{\mbox{\isa{thin{\isaliteral{5F}{\isacharunderscore}}tac}}}~\isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} deletes the specified assumption - from a subgoal; note that \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} may contain schematic variables. - See also \verb|thin_tac| in \cite{isabelle-implementation}. + \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{5C3C7068693E}{\isasymphi}}} adds \isa{{\isaliteral{5C3C7068693E}{\isasymphi}}} as an - assumption to a subgoal. See also \verb|subgoal_tac| and \verb|subgoals_tac| in \cite{isabelle-implementation}. + \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 assumptions of a - goal by \isa{n} positions: from right to left if \isa{n} is + \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. See also \verb|rotate_tac| in - \cite{isabelle-implementation}. + 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 diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/HOL_Specific.tex --- a/doc-src/IsarRef/Thy/document/HOL_Specific.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/HOL_Specific.tex Tue Feb 14 12:40:55 2012 +0100 @@ -83,11 +83,19 @@ \isamarkuptrue% % \begin{isamarkuptext}% -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. +\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: @@ -109,14 +117,6 @@ introduction rule. The default rule declarations of Isabelle/HOL already take care of most common situations. - \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} - \begin{railoutput} \rail@begin{10}{} \rail@bar @@ -784,10 +784,10 @@ 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. + 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}% @@ -851,10 +851,11 @@ 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}}}. + 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 @@ -863,6 +864,7 @@ 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 @@ -881,13 +883,13 @@ \isamarkuptrue% % \begin{isamarkuptext}% -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{matharray}{rcl} +\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}}}}}[] @@ -1487,8 +1489,12 @@ \isamarkuptrue% % \begin{isamarkuptext}% -A Gordon/HOL-style type definition is a certain axiom scheme - that identifies a new type with a subset of an existing type. More +\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 @@ -1506,10 +1512,6 @@ 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{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} - \begin{railoutput} \rail@begin{2}{} \rail@term{\hyperlink{command.HOL.typedef}{\mbox{\isa{\isacommand{typedef}}}}}[] @@ -1768,13 +1770,7 @@ \isamarkuptrue% % \begin{isamarkuptext}% -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{matharray}{rcl} +\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\_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}}}\\ @@ -1795,6 +1791,12 @@ \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}}}}}[] @@ -1882,16 +1884,18 @@ \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 + \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-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}} prints quotient map functions. + \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-quotmaps}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotmaps}}}} prints quotient map + functions. \item \hyperlink{command.HOL.print-quotients}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}quotients}}}} prints quotients. @@ -1967,6 +1971,11 @@ \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}}}}[] @@ -1988,35 +1997,28 @@ \end{railoutput} - 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{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 nullary type constructors. 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 + 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. - + 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. @@ -2036,18 +2038,22 @@ \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} - The \hyperlink{method.HOL.arith}{\mbox{\isa{arith}}} method decides linear arithmetic problems - (on types \isa{nat}, \isa{int}, \isa{real}). Any current - facts are inserted into the goal before running the procedure. - - The \hyperlink{attribute.HOL.arith}{\mbox{\isa{arith}}} attribute declares facts that are - always supplied to the arithmetic provers implicitly. - - The \hyperlink{attribute.HOL.arith-split}{\mbox{\isa{arith{\isaliteral{5F}{\isacharunderscore}}split}}} attribute declares case split + \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. - Note that a simpler (but faster) arithmetic prover is - already invoked by the Simplifier.% + \end{description} + + Note that a simpler (but faster) arithmetic prover is already + invoked by the Simplifier.% \end{isamarkuptext}% \isamarkuptrue% % @@ -2071,10 +2077,12 @@ \end{railoutput} - The \hyperlink{method.HOL.iprover}{\mbox{\isa{iprover}}} method 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. + \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 @@ -2083,7 +2091,9 @@ 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.% + number of rule premises will be taken into account here. + + \end{description}% \end{isamarkuptext}% \isamarkuptrue% % @@ -2129,15 +2139,19 @@ \end{railoutput} - The \hyperlink{method.HOL.meson}{\mbox{\isa{meson}}} method implements Loveland's model elimination - procedure \cite{loveland-78}. See \verb|~~/src/HOL/ex/Meson_Test.thy| for - examples. - - The \hyperlink{method.HOL.metis}{\mbox{\isa{metis}}} method 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 \verb|~~/src/HOL/Metis_Examples| directory contains several small theories - developed to a large extent using Metis.% + \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% % @@ -2161,11 +2175,14 @@ \end{railoutput} - The \hyperlink{method.HOL.coherent}{\mbox{\isa{coherent}}} method 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.% + \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% % @@ -2274,27 +2291,24 @@ \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.try-methods}{\mbox{\isa{\isacommand{try{\isaliteral{5F}{\isacharunderscore}}methods}}}} attempts to prove a subgoal using a combination - of standard proof methods (\isa{auto}, \isa{simp}, \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.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.try-methods}{\mbox{\isa{\isacommand{try{\isaliteral{5F}{\isacharunderscore}}methods}}}} 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 (\isa{{\isaliteral{22}{\isachardoublequote}}solve{\isaliteral{5F}{\isacharunderscore}}direct{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}quickcheck{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}try{\isaliteral{5F}{\isacharunderscore}}methods{\isaliteral{22}{\isachardoublequote}}}, \isa{{\isaliteral{22}{\isachardoublequote}}sledgehammer{\isaliteral{22}{\isachardoublequote}}}, - \isa{{\isaliteral{22}{\isachardoublequote}}nitpick{\isaliteral{22}{\isachardoublequote}}}). - - \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. + 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.try-methods}{\mbox{\isa{\isacommand{try{\isaliteral{5F}{\isacharunderscore}}methods}}}}, \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}% @@ -2415,54 +2429,51 @@ \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}. + 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: + 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 - \isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}exhaustive{\isaliteral{5F}{\isacharunderscore}}active}, \isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}random{\isaliteral{5F}{\isacharunderscore}}active}, - \isa{quickcheck{\isaliteral{5F}{\isacharunderscore}}narrowing{\isaliteral{5F}{\isacharunderscore}}active} are set to true. + 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. + counterexample, but not potentially spurious counterexamples due + to underspecified functions. \item[\isa{eval}] takes a term or a list of terms and evaluates - these terms under the variable assignment found by quickcheck. + these terms under the variable assignment found by quickcheck. \item[\isa{iterations}] sets how many sets of assignments are generated for each particular size. @@ -2481,39 +2492,38 @@ \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{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. + 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. + 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: + 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 $\infty$. - - \item[\isa{maxvars}] specifies the maximum number of Boolean variables - to use when transforming the term into a propositional formula. - Nonpositive values mean $\infty$. + \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. @@ -2523,22 +2533,19 @@ \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}). + 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. + 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. \end{description}% \end{isamarkuptext}% diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Inner_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Inner_Syntax.tex Tue Feb 14 12:40:55 2012 +0100 @@ -22,11 +22,35 @@ } \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 + translations} --- either as rewrite systems on first-order ASTs + (\secref{sec:syn-trans}) or ML functions on ASTs or \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C616D6264613E}{\isasymlambda}}{\isaliteral{22}{\isachardoublequote}}}-terms that represent parse trees (\secref{sec:tr-funs}).% +\end{isamarkuptext}% +\isamarkuptrue% +% \isamarkupsection{Printing logical entities% } \isamarkuptrue% % -\isamarkupsubsection{Diagnostic commands% +\isamarkupsubsection{Diagnostic commands \label{sec:print-diag}% } \isamarkuptrue% % @@ -152,10 +176,11 @@ All of the diagnostic commands above admit a list of \isa{modes} to be specified, which is appended to the current print mode; see - \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 + 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 @@ -277,6 +302,79 @@ \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% @@ -307,65 +405,71 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{Mixfix annotations% +\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 specifications etc.\ also admit mixfix - annotations. + 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{3}{\indexdef{}{syntax}{infix}\hypertarget{syntax.infix}{\hyperlink{syntax.infix}{\mbox{\isa{infix}}}}} +\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@term{\isa{\isakeyword{infix}}}[] +\rail@nont{\isa{mfix}}[] \rail@nextbar{1} -\rail@term{\isa{\isakeyword{infixl}}}[] -\rail@nextbar{2} -\rail@term{\isa{\isakeyword{infixr}}}[] +\rail@term{\isa{\isakeyword{structure}}}[] \rail@endbar -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] \rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] \rail@end -\rail@begin{5}{\indexdef{}{syntax}{mixfix}\hypertarget{syntax.mixfix}{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}} +\rail@begin{7}{\isa{mfix}} \rail@bar -\rail@nont{\hyperlink{syntax.infix}{\mbox{\isa{infix}}}}[] +\rail@nont{\hyperlink{syntax.template}{\mbox{\isa{template}}}}[] +\rail@bar \rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@nont{\hyperlink{syntax.string}{\mbox{\isa{string}}}}[] -\rail@bar -\rail@nextbar{2} \rail@nont{\isa{prios}}[] \rail@endbar \rail@bar -\rail@nextbar{2} +\rail@nextbar{1} \rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] \rail@endbar -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] +\rail@nextbar{2} +\rail@bar +\rail@term{\isa{\isakeyword{infix}}}[] \rail@nextbar{3} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] +\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.string}{\mbox{\isa{string}}}}[] +\rail@nont{\hyperlink{syntax.template}{\mbox{\isa{template}}}}[] \rail@bar -\rail@nextbar{4} +\rail@nextbar{6} \rail@nont{\isa{prios}}[] \rail@endbar \rail@nont{\hyperlink{syntax.nat}{\mbox{\isa{nat}}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] \rail@endbar \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@bar -\rail@nont{\hyperlink{syntax.mixfix}{\mbox{\isa{mixfix}}}}[] -\rail@nextbar{1} -\rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[] -\rail@term{\isa{\isakeyword{structure}}}[] -\rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[] -\rail@endbar +\rail@begin{1}{\isa{template}} +\rail@nont{\isa{string}}[] \rail@end \rail@begin{2}{\isa{prios}} \rail@term{\isa{{\isaliteral{5B}{\isacharbrackleft}}}}[] @@ -379,28 +483,36 @@ \end{railoutput} - Here the \hyperlink{syntax.string}{\mbox{\isa{string}}} specifications refer to the actual mixfix - template, which 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. - - \medskip 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. + 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. + 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 @@ -464,17 +576,87 @@ \end{description} - For example, the template \verb|(_ +/ _)| specifies an infix - operator. There are two argument positions; the delimiter - \verb|+| is preceded by a space and followed by a space or - line break; the entire phrase is a pretty printing block. + 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} - The general idea of pretty printing with blocks and breaks is also - described in \cite{paulson-ml2}.% + \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% % -\isamarkupsection{Explicit notation% +\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% % @@ -487,6 +669,13 @@ \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 @@ -553,14 +742,14 @@ \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. @@ -568,11 +757,7 @@ \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} - - Note that the more primitive commands \hyperlink{command.syntax}{\mbox{\isa{\isacommand{syntax}}}} and - \hyperlink{command.no-syntax}{\mbox{\isa{\isacommand{no{\isaliteral{5F}{\isacharunderscore}}syntax}}}} (\secref{sec:syn-trans}) provide raw access - to the syntax tables of a global theory.% + \end{description}% \end{isamarkuptext}% \isamarkuptrue% % @@ -580,6 +765,59 @@ } \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}{xstr}\hypertarget{syntax.inner.xstr}{\hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}}} & = & \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.xstr}{\mbox{\isa{xstr}}} 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% @@ -655,15 +893,13 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsubsection{The Pure grammar% +\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 as follows: - - %FIXME syntax for "index" (?) - %FIXME "op" versions of ==> etc. (?) +The priority grammar of the \isa{{\isaliteral{22}{\isachardoublequote}}Pure{\isaliteral{22}{\isachardoublequote}}} theory is defined + approximately like this: \begin{center} \begin{supertabular}{rclr} @@ -689,21 +925,28 @@ \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} \\\\ @@ -713,16 +956,17 @@ \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}}id\ \ {\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}}\ id\ \ {\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{id} \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \isa{{\isaliteral{22}{\isachardoublequote}}longid\ \ {\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}}\ longid{\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{longid} \\ + & \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}}} \\\\ + & \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}}}} & = & \isa{{\isaliteral{22}{\isachardoublequote}}id\ \ {\isaliteral{7C}{\isacharbar}}\ \ longid\ \ {\isaliteral{7C}{\isacharbar}}\ \ {\isaliteral{22}{\isachardoublequote}}}\verb|{}| \\ - & \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{7C}{\isacharbar}}{\isaliteral{22}{\isachardoublequote}}} & \verb|{| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}id\ {\isaliteral{7C}{\isacharbar}}\ longid{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|,| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{28}{\isacharparenleft}}id\ {\isaliteral{7C}{\isacharbar}}\ longid{\isaliteral{29}{\isacharparenright}}{\isaliteral{22}{\isachardoublequote}}} \verb|}| \\ + \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} @@ -759,6 +1003,10 @@ (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. @@ -827,64 +1075,121 @@ 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% % -\isamarkupsection{Lexical matters \label{sec:inner-lex}% +\isamarkupsubsection{Inspecting the syntax% } \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{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{enumerate} + \begin{description} - \item \emph{delimiters} --- the literal tokens occurring in - productions of the given priority grammar (cf.\ - \secref{sec:priority-grammar}); + \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. - \item \emph{named tokens} --- various categories of identifiers etc. + \begin{description} - \end{enumerate} + \item \isa{{\isaliteral{22}{\isachardoublequote}}lexicon{\isaliteral{22}{\isachardoublequote}}} lists the delimiters of the inner token + language; see \secref{sec:inner-lex}. - 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. + \item \isa{{\isaliteral{22}{\isachardoublequote}}prods{\isaliteral{22}{\isachardoublequote}}} lists the productions of the underlying + priority grammar; see \secref{sec:priority-grammar}. - \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}). + 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. - \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}}} \\ + 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. - \indexdef{inner}{syntax}{xstr}\hypertarget{syntax.inner.xstr}{\hyperlink{syntax.inner.xstr}{\mbox{\isa{xstr}}}} & = & \verb|''| \isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C646F74733E}{\isasymdots}}{\isaliteral{22}{\isachardoublequote}}} \verb|''| \\ - \end{supertabular} - \end{center} + \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}. - 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.xstr}{\mbox{\isa{xstr}}} 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). + \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}. - 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.% + \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% % -\isamarkupsection{Syntax and translations \label{sec:syn-trans}% +\isamarkupsubsection{Ambiguity of parsed expressions% +} +\isamarkuptrue% +% +\begin{isamarkuptext}% +\begin{tabular}{rcll} + \indexdef{}{attribute}{syntax\_ambiguity\_level}\hypertarget{attribute.syntax-ambiguity-level}{\hyperlink{attribute.syntax-ambiguity-level}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}level}}}} & : & \isa{attribute} & default \isa{{\isadigit{1}}} \\ + \end{tabular} + + \begin{mldecls} + \indexdef{}{ML}{Syntax.ambiguity\_limit}\verb|Syntax.ambiguity_limit: int Config.T| \\ %FIXME attribute + \end{mldecls} + + 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-level}{\mbox{\isa{syntax{\isaliteral{5F}{\isacharunderscore}}ambiguity{\isaliteral{5F}{\isacharunderscore}}level}}} determines the number of + parser results that are tolerated without printing a detailed + message. + + \item \verb|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}% +\end{isamarkuptext}% +\isamarkuptrue% +% +\isamarkupsection{Raw syntax and translations \label{sec:syn-trans}% } \isamarkuptrue% % @@ -897,6 +1202,15 @@ \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}}} \\ \end{matharray} + Unlike mixfix notation for existing formal entities + (\secref{sec:notation}), raw syntax declarations provide full access + to the priority grammar of the inner syntax. 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}}}}}[] @@ -971,34 +1285,87 @@ \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}}\ decls{\isaliteral{22}{\isachardoublequote}}} is similar to - \hyperlink{command.consts}{\mbox{\isa{\isacommand{consts}}}}~\isa{decls}, except that the actual logical - signature extension is omitted. Thus the context free grammar of - Isabelle's inner syntax may be augmented in arbitrary ways, - independently of the logic. The \isa{mode} argument refers to the - print mode that the grammar rules belong; unless the \indexref{}{keyword}{output}\hyperlink{keyword.output}{\mbox{\isa{\isakeyword{output}}}} indicator is given, all productions are added both to the - input and output grammar. - + \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 unrelated 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): parse~/ print rules (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C72696768746C656674686172706F6F6E733E}{\isasymrightleftharpoons}}{\isaliteral{22}{\isachardoublequote}}}), parse rules (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C7269676874686172706F6F6E75703E}{\isasymrightharpoonup}}{\isaliteral{22}{\isachardoublequote}}}), or print rules (\isa{{\isaliteral{22}{\isachardoublequote}}{\isaliteral{5C3C6C656674686172706F6F6E646F776E3E}{\isasymleftharpoondown}}{\isaliteral{22}{\isachardoublequote}}}). Translation patterns may be prefixed by the syntactic category to be used for parsing; the default is \isa{logic}. - + \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. - \end{description}% + \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 \hyperlink{theory.List}{\mbox{\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 \hyperlink{theory.List}{\mbox{\isa{List}}} in Isabelle/HOL. + + \end{itemize}% \end{isamarkuptext}% \isamarkuptrue% % @@ -1059,9 +1426,6 @@ translations functions may refer to specific theory declarations or auxiliary proof data. - See also \cite{isabelle-ref} for more information on the general - concept of syntax transformations in Isabelle. - %FIXME proper antiquotations \begin{ttbox} val parse_ast_translation: @@ -1074,67 +1438,11 @@ (string * (Proof.context -> typ -> term list -> term)) list val print_ast_translation: (string * (Proof.context -> ast list -> ast)) list -\end{ttbox}% -\end{isamarkuptext}% -\isamarkuptrue% -% -\isamarkupsection{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. +\end{ttbox} - 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}% + \medskip See also the chapter on ``Syntax Transformations'' in old + \cite{isabelle-ref} for further details on translations on parse + trees.% \end{isamarkuptext}% \isamarkuptrue% % diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/ML_Tactic.tex --- a/doc-src/IsarRef/Thy/document/ML_Tactic.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/ML_Tactic.tex Tue Feb 14 12:40:55 2012 +0100 @@ -125,10 +125,9 @@ \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.\ (see \cite{isabelle-ref}). 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}).% +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% % @@ -181,23 +180,25 @@ \end{tabular} \medskip - \medskip \verb|CHANGED| (see \cite{isabelle-ref}) 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|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-ref}) 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 + \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, such as solving - goals in a different order. + this usually has a different operational behavior: subgoals are + solved in a different order. - \medskip Iterated resolution, such as \verb|REPEAT (FIRSTGOAL|\isasep\isanewline% -\verb| (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}).% + \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% % diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Misc.tex --- a/doc-src/IsarRef/Thy/document/Misc.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Misc.tex Tue Feb 14 12:40:55 2012 +0100 @@ -217,36 +217,6 @@ \end{isamarkuptext}% \isamarkuptrue% % -\isamarkupsection{History commands \label{sec:history}% -} -\isamarkuptrue% -% -\begin{isamarkuptext}% -\begin{matharray}{rcl} - \indexdef{}{command}{undo}\hypertarget{command.undo}{\hyperlink{command.undo}{\mbox{\isa{\isacommand{undo}}}}}^{{ * }{ * }} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ any{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{linear\_undo}\hypertarget{command.linear-undo}{\hyperlink{command.linear-undo}{\mbox{\isa{\isacommand{linear{\isaliteral{5F}{\isacharunderscore}}undo}}}}}^{{ * }{ * }} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ any{\isaliteral{22}{\isachardoublequote}}} \\ - \indexdef{}{command}{kill}\hypertarget{command.kill}{\hyperlink{command.kill}{\mbox{\isa{\isacommand{kill}}}}}^{{ * }{ * }} & : & \isa{{\isaliteral{22}{\isachardoublequote}}any\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ any{\isaliteral{22}{\isachardoublequote}}} \\ - \end{matharray} - - The Isabelle/Isar top-level maintains a two-stage history, for - theory and proof state transformation. Basically, any command can - be undone using \hyperlink{command.undo}{\mbox{\isa{\isacommand{undo}}}}, excluding mere diagnostic - elements. Note that a theorem statement with a \emph{finished} - proof is treated as a single unit by \hyperlink{command.undo}{\mbox{\isa{\isacommand{undo}}}}. In - contrast, the variant \hyperlink{command.linear-undo}{\mbox{\isa{\isacommand{linear{\isaliteral{5F}{\isacharunderscore}}undo}}}} admits to step back - into the middle of a proof. The \hyperlink{command.kill}{\mbox{\isa{\isacommand{kill}}}} command aborts - the current history node altogether, discontinuing a proof or even - the whole theory. This operation is \emph{not} undo-able. - - \begin{warn} - History commands should never be used with user interfaces such as - Proof~General \cite{proofgeneral,Aspinall:TACAS:2000}, which takes - care of stepping forth and back itself. Interfering by manual - \hyperlink{command.undo}{\mbox{\isa{\isacommand{undo}}}}, \hyperlink{command.linear-undo}{\mbox{\isa{\isacommand{linear{\isaliteral{5F}{\isacharunderscore}}undo}}}}, or even \hyperlink{command.kill}{\mbox{\isa{\isacommand{kill}}}} commands would quickly result in utter confusion. - \end{warn}% -\end{isamarkuptext}% -\isamarkuptrue% -% \isamarkupsection{System commands% } \isamarkuptrue% diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Outer_Syntax.tex --- a/doc-src/IsarRef/Thy/document/Outer_Syntax.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Outer_Syntax.tex Tue Feb 14 12:40:55 2012 +0100 @@ -18,7 +18,7 @@ % \endisadelimtheory % -\isamarkupchapter{Outer syntax% +\isamarkupchapter{Outer syntax --- the theory language% } \isamarkuptrue% % diff -r ec2e20b27638 -r 915af80f74b3 doc-src/IsarRef/Thy/document/Proof.tex --- a/doc-src/IsarRef/Thy/document/Proof.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/IsarRef/Thy/document/Proof.tex Tue Feb 14 12:40:55 2012 +0100 @@ -167,11 +167,7 @@ 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}}}}''. - - \medskip The \hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}} command is undo-able, unlike - \indexref{}{command}{kill}\hyperlink{command.kill}{\mbox{\isa{\isacommand{kill}}}} (see \secref{sec:history}). The effect is to - get back to the theory just before the opening of the proof.% + the keyword ``\hyperlink{command.oops}{\mbox{\isa{\isacommand{oops}}}}''.% \end{isamarkuptext}% \isamarkuptrue% % @@ -1113,7 +1109,7 @@ \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 parallel). This corresponds to the \verb|op MRS| operation in + (in parallel). This corresponds to the \verb|MRS| operation in ML, but note the reversed order. Positions may be effectively skipped by including ``\isa{{\isaliteral{5F}{\isacharunderscore}}}'' (underscore) as argument. diff -r ec2e20b27638 -r 915af80f74b3 doc-src/Ref/Makefile --- a/doc-src/Ref/Makefile Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/Ref/Makefile Tue Feb 14 12:40:55 2012 +0100 @@ -9,7 +9,7 @@ include ../Makefile.in NAME = ref -FILES = ref.tex tactic.tex tctical.tex thm.tex defining.tex syntax.tex \ +FILES = ref.tex tactic.tex thm.tex syntax.tex \ substitution.tex simplifier.tex classical.tex ../proof.sty \ ../iman.sty ../extra.sty ../ttbox.sty ../manual.bib diff -r ec2e20b27638 -r 915af80f74b3 doc-src/Ref/ref.tex --- a/doc-src/Ref/ref.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/Ref/ref.tex Tue Feb 14 12:40:55 2012 +0100 @@ -48,9 +48,7 @@ \pagenumbering{roman} \tableofcontents \clearfirst \include{tactic} -\include{tctical} \include{thm} -\include{defining} \include{syntax} \include{substitution} \include{simplifier} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/Ref/tactic.tex --- a/doc-src/Ref/tactic.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/Ref/tactic.tex Tue Feb 14 12:40:55 2012 +0100 @@ -9,8 +9,6 @@ \begin{ttbox} cut_facts_tac : thm list -> int -> tactic cut_inst_tac : (string*string)list -> thm -> int -> tactic -subgoal_tac : string -> int -> tactic -subgoals_tac : string list -> int -> tactic \end{ttbox} These tactics add assumptions to a subgoal. \begin{ttdescription} @@ -28,155 +26,6 @@ described in {\S}\ref{res_inst_tac}. It adds the resulting theorem as a new assumption to subgoal~$i$. -\item[\ttindexbold{subgoal_tac} {\it formula} {\it i}] -adds the {\it formula} as an assumption to subgoal~$i$, and inserts the same -{\it formula} as a new subgoal, $i+1$. - -\item[\ttindexbold{subgoals_tac} {\it formulae} {\it i}] - uses {\tt subgoal_tac} to add the members of the list of {\it - formulae} as assumptions to subgoal~$i$. -\end{ttdescription} - - -\subsection{``Putting off'' a subgoal} -\begin{ttbox} -defer_tac : int -> tactic -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{defer_tac} {\it i}] - moves subgoal~$i$ to the last position in the proof state. It can be - useful when correcting a proof script: if the tactic given for subgoal~$i$ - fails, calling {\tt defer_tac} instead will let you continue with the rest - of the script. - - The tactic fails if subgoal~$i$ does not exist or if the proof state - contains type unknowns. -\end{ttdescription} - - -\subsection{Definitions and meta-level rewriting} \label{sec:rewrite_goals} -\index{tactics!meta-rewriting|bold}\index{meta-rewriting|bold} -\index{definitions} - -Definitions in Isabelle have the form $t\equiv u$, where $t$ is typically a -constant or a constant applied to a list of variables, for example $\it -sqr(n)\equiv n\times n$. Conditional definitions, $\phi\Imp t\equiv u$, -are also supported. {\bf Unfolding} the definition ${t\equiv u}$ means using -it as a rewrite rule, replacing~$t$ by~$u$ throughout a theorem. {\bf -Folding} $t\equiv u$ means replacing~$u$ by~$t$. Rewriting continues until -no rewrites are applicable to any subterm. - -There are rules for unfolding and folding definitions; Isabelle does not do -this automatically. The corresponding tactics rewrite the proof state, -yielding a single next state. See also the {\tt goalw} command, which is the -easiest way of handling definitions. -\begin{ttbox} -rewrite_goals_tac : thm list -> tactic -rewrite_tac : thm list -> tactic -fold_goals_tac : thm list -> tactic -fold_tac : thm list -> tactic -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{rewrite_goals_tac} {\it defs}] -unfolds the {\it defs} throughout the subgoals of the proof state, while -leaving the main goal unchanged. Use \ttindex{SELECT_GOAL} to restrict it to a -particular subgoal. - -\item[\ttindexbold{rewrite_tac} {\it defs}] -unfolds the {\it defs} throughout the proof state, including the main goal ---- not normally desirable! - -\item[\ttindexbold{fold_goals_tac} {\it defs}] -folds the {\it defs} throughout the subgoals of the proof state, while -leaving the main goal unchanged. - -\item[\ttindexbold{fold_tac} {\it defs}] -folds the {\it defs} throughout the proof state. -\end{ttdescription} - -\begin{warn} - These tactics only cope with definitions expressed as meta-level - equalities ($\equiv$). More general equivalences are handled by the - simplifier, provided that it is set up appropriately for your logic - (see Chapter~\ref{chap:simplification}). -\end{warn} - -\subsection{Theorems useful with tactics} -\index{theorems!of pure theory} -\begin{ttbox} -asm_rl: thm -cut_rl: thm -\end{ttbox} -\begin{ttdescription} -\item[\tdx{asm_rl}] -is $\psi\Imp\psi$. Under elim-resolution it does proof by assumption, and -\hbox{\tt eresolve_tac (asm_rl::{\it thms}) {\it i}} is equivalent to -\begin{ttbox} -assume_tac {\it i} ORELSE eresolve_tac {\it thms} {\it i} -\end{ttbox} - -\item[\tdx{cut_rl}] -is $\List{\psi\Imp\theta,\psi}\Imp\theta$. It is useful for inserting -assumptions; it underlies {\tt forward_tac}, {\tt cut_facts_tac} -and {\tt subgoal_tac}. -\end{ttdescription} - - -\section{Obscure tactics} - -\subsection{Manipulating assumptions} -\index{assumptions!rotating} -\begin{ttbox} -thin_tac : string -> int -> tactic -rotate_tac : int -> int -> tactic -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{thin_tac} {\it formula} $i$] -\index{assumptions!deleting} -deletes the specified assumption from subgoal $i$. Often the assumption -can be abbreviated, replacing subformul{\ae} by unknowns; the first matching -assumption will be deleted. Removing useless assumptions from a subgoal -increases its readability and can make search tactics run faster. - -\item[\ttindexbold{rotate_tac} $n$ $i$] -\index{assumptions!rotating} -rotates the assumptions of subgoal $i$ by $n$ positions: from right to left -if $n$ is positive, and from left to right if $n$ is negative. This is -sometimes necessary in connection with \ttindex{asm_full_simp_tac}, which -processes assumptions from left to right. -\end{ttdescription} - - -\subsection{Tidying the proof state} -\index{duplicate subgoals!removing} -\index{parameters!removing unused} -\index{flex-flex constraints} -\begin{ttbox} -distinct_subgoals_tac : tactic -prune_params_tac : tactic -flexflex_tac : tactic -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{distinct_subgoals_tac}] removes duplicate subgoals from a - proof state. (These arise especially in ZF, where the subgoals are - essentially type constraints.) - -\item[\ttindexbold{prune_params_tac}] - removes unused parameters from all subgoals of the proof state. It works - by rewriting with the theorem $(\Forall x. V)\equiv V$. This tactic can - make the proof state more readable. It is used with - \ttindex{rule_by_tactic} to simplify the resulting theorem. - -\item[\ttindexbold{flexflex_tac}] - removes all flex-flex pairs from the proof state by applying the trivial - unifier. This drastic step loses information, and should only be done as - the last step of a proof. - - Flex-flex constraints arise from difficult cases of higher-order - unification. To prevent this, use \ttindex{res_inst_tac} to instantiate - some variables in a rule~({\S}\ref{res_inst_tac}). Normally flex-flex - constraints can be ignored; they often disappear as unknowns get - instantiated. \end{ttdescription} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/Ref/thm.tex --- a/doc-src/Ref/thm.tex Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/Ref/thm.tex Tue Feb 14 12:40:55 2012 +0100 @@ -11,57 +11,6 @@ \section{Basic operations on theorems} -\subsection{Forward proof: joining rules by resolution} -\index{theorems!joining by resolution} -\index{resolution}\index{forward proof} -\begin{ttbox} -RSN : thm * (int * thm) -> thm \hfill\textbf{infix} -RS : thm * thm -> thm \hfill\textbf{infix} -MRS : thm list * thm -> thm \hfill\textbf{infix} -OF : thm * thm list -> thm \hfill\textbf{infix} -RLN : thm list * (int * thm list) -> thm list \hfill\textbf{infix} -RL : thm list * thm list -> thm list \hfill\textbf{infix} -MRL : thm list list * thm list -> thm list \hfill\textbf{infix} -\end{ttbox} -Joining rules together is a simple way of deriving new rules. These -functions are especially useful with destruction rules. To store -the result in the theorem database, use \ttindex{bind_thm} -(\S\ref{ExtractingAndStoringTheProvedTheorem}). -\begin{ttdescription} -\item[\tt$thm@1$ RSN $(i,thm@2)$] \indexbold{*RSN} - resolves the conclusion of $thm@1$ with the $i$th premise of~$thm@2$. - Unless there is precisely one resolvent it raises exception - \xdx{THM}; in that case, use {\tt RLN}. - -\item[\tt$thm@1$ RS $thm@2$] \indexbold{*RS} -abbreviates \hbox{\tt$thm@1$ RSN $(1,thm@2)$}. Thus, it resolves the -conclusion of $thm@1$ with the first premise of~$thm@2$. - -\item[\tt {$[thm@1,\ldots,thm@n]$} MRS $thm$] \indexbold{*MRS} - uses {\tt RSN} to resolve $thm@i$ against premise~$i$ of $thm$, for - $i=n$, \ldots,~1. This applies $thm@n$, \ldots, $thm@1$ to the first $n$ - premises of $thm$. Because the theorems are used from right to left, it - does not matter if the $thm@i$ create new premises. {\tt MRS} is useful - for expressing proof trees. - -\item[\tt {$thm$ OF $[thm@1,\ldots,thm@n]$}] \indexbold{*OF} is the same as - \texttt{$[thm@1,\ldots,thm@n]$ MRS $thm$}, with slightly more readable - argument order, though. - -\item[\tt$thms@1$ RLN $(i,thms@2)$] \indexbold{*RLN} - joins lists of theorems. For every $thm@1$ in $thms@1$ and $thm@2$ in - $thms@2$, it resolves the conclusion of $thm@1$ with the $i$th premise - of~$thm@2$, accumulating the results. - -\item[\tt$thms@1$ RL $thms@2$] \indexbold{*RL} -abbreviates \hbox{\tt$thms@1$ RLN $(1,thms@2)$}. - -\item[\tt {$[thms@1,\ldots,thms@n]$} MRL $thms$] \indexbold{*MRL} -is analogous to {\tt MRS}, but combines theorem lists rather than theorems. -It too is useful for expressing proof trees. -\end{ttdescription} - - \subsection{Expanding definitions in theorems} \index{meta-rewriting!in theorems} \begin{ttbox} @@ -193,13 +142,7 @@ cprems_of : thm -> cterm list nprems_of : thm -> int tpairs_of : thm -> (term*term) list -sign_of_thm : thm -> Sign.sg theory_of_thm : thm -> theory -dest_state : thm * int -> (term*term) list * term list * term * term -rep_thm : thm -> \{sign_ref: Sign.sg_ref, der: bool * deriv, maxidx: int, - shyps: sort list, hyps: term list, prop: term\} -crep_thm : thm -> \{sign_ref: Sign.sg_ref, der: bool * deriv, maxidx: int, - shyps: sort list, hyps: cterm list, prop:{\ts}cterm\} \end{ttbox} \begin{ttdescription} \item[\ttindexbold{cprop_of} $thm$] returns the statement of $thm$ as @@ -221,28 +164,10 @@ \item[\ttindexbold{tpairs_of} $thm$] returns the flex-flex constraints of $thm$. -\item[\ttindexbold{sign_of_thm} $thm$] returns the signature - associated with $thm$. - \item[\ttindexbold{theory_of_thm} $thm$] returns the theory associated with $thm$. Note that this does a lookup in Isabelle's global database of loaded theories. -\item[\ttindexbold{dest_state} $(thm,i)$] -decomposes $thm$ as a tuple containing a list of flex-flex constraints, a -list of the subgoals~1 to~$i-1$, subgoal~$i$, and the rest of the theorem -(this will be an implication if there are more than $i$ subgoals). - -\item[\ttindexbold{rep_thm} $thm$] decomposes $thm$ as a record - containing the statement of~$thm$ ({\tt prop}), its list of - meta-assumptions ({\tt hyps}), its derivation ({\tt der}), a bound - on the maximum subscript of its unknowns ({\tt maxidx}), and a - reference to its signature ({\tt sign_ref}). The {\tt shyps} field - is discussed below. - -\item[\ttindexbold{crep_thm} $thm$] like \texttt{rep_thm}, but returns - the hypotheses and statement as certified terms. - \end{ttdescription} @@ -403,21 +328,6 @@ Most of these rules have the sole purpose of implementing particular tactics. There are few occasions for applying them directly to a theorem. -\subsection{Proof by assumption} -\index{meta-assumptions} -\begin{ttbox} -assumption : int -> thm -> thm Seq.seq -eq_assumption : int -> thm -> thm -\end{ttbox} -\begin{ttdescription} -\item[\ttindexbold{assumption} {\it i} $thm$] -attempts to solve premise~$i$ of~$thm$ by assumption. - -\item[\ttindexbold{eq_assumption}] -is like {\tt assumption} but does not use unification. -\end{ttdescription} - - \subsection{Resolution} \index{resolution} \begin{ttbox} @@ -475,20 +385,9 @@ \subsection{Other meta-rules} \begin{ttbox} -trivial : cterm -> thm -lift_rule : (thm * int) -> thm -> thm rename_params_rule : string list * int -> thm -> thm -flexflex_rule : thm -> thm Seq.seq \end{ttbox} \begin{ttdescription} -\item[\ttindexbold{trivial} $ct$] -makes the theorem \(\phi\Imp\phi\), where $\phi$ is the value of~$ct$. -This is the initial state for a goal-directed proof of~$\phi$. The rule -checks that $ct$ has type~$prop$. - -\item[\ttindexbold{lift_rule} ($state$, $i$) $rule$] \index{lifting} -prepares $rule$ for resolution by lifting it over the parameters and -assumptions of subgoal~$i$ of~$state$. \item[\ttindexbold{rename_params_rule} ({\it names}, {\it i}) $thm$] uses the $names$ to rename the parameters of premise~$i$ of $thm$. The @@ -497,8 +396,6 @@ ensure that all the parameters are distinct. \index{parameters!renaming} -\item[\ttindexbold{flexflex_rule} $thm$] \index{flex-flex constraints} -removes all flex-flex pairs from $thm$ using the trivial unifier. \end{ttdescription} \index{meta-rules|)} diff -r ec2e20b27638 -r 915af80f74b3 doc-src/antiquote_setup.ML --- a/doc-src/antiquote_setup.ML Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/antiquote_setup.ML Tue Feb 14 12:40:55 2012 +0100 @@ -52,6 +52,9 @@ fun ml_val (txt1, "") = "fn _ => (" ^ txt1 ^ ");" | ml_val (txt1, txt2) = "fn _ => (" ^ txt1 ^ " : " ^ txt2 ^ ");"; +fun ml_op (txt1, "") = "fn _ => (op " ^ txt1 ^ ");" + | ml_op (txt1, txt2) = "fn _ => (op " ^ txt1 ^ " : " ^ txt2 ^ ");"; + fun ml_type (txt1, "") = "val _ = NONE : (" ^ txt1 ^ ") option;" | ml_type (txt1, txt2) = "val _ = [NONE : (" ^ txt1 ^ ") option, NONE : (" ^ txt2 ^ ") option];"; @@ -95,6 +98,7 @@ val index_ml_setup = index_ml @{binding index_ML} "" ml_val #> + index_ml @{binding index_ML_op} "infix" ml_op #> index_ml @{binding index_ML_type} "type" ml_type #> index_ml @{binding index_ML_exn} "exception" ml_exn #> index_ml @{binding index_ML_structure} "structure" ml_structure #> diff -r ec2e20b27638 -r 915af80f74b3 doc-src/manual.bib --- a/doc-src/manual.bib Sat Feb 11 13:41:36 2012 +0100 +++ b/doc-src/manual.bib Tue Feb 14 12:40:55 2012 +0100 @@ -1135,6 +1135,14 @@ number = {IC/2004/64} } +@Article{Oppen:1980, + author = {D. C. Oppen}, + title = {Pretty Printing}, + journal = {ACM Transactions on Programming Languages and Systems}, + year = 1980, + volume = 2, + number = 4} + @Manual{pvs-language, title = {The {PVS} specification language}, author = {S. Owre and N. Shankar and J. M. Rushby}, diff -r ec2e20b27638 -r 915af80f74b3 src/Pure/Thy/thy_output.ML --- a/src/Pure/Thy/thy_output.ML Sat Feb 11 13:41:36 2012 +0100 +++ b/src/Pure/Thy/thy_output.ML Tue Feb 14 12:40:55 2012 +0100 @@ -670,6 +670,7 @@ val _ = Context.>> (Context.map_theory (ml_text (Binding.name "ML") (ml_enclose "fn _ => (" ");") #> + ml_text (Binding.name "ML_op") (ml_enclose "fn _ => (op " ");") #> ml_text (Binding.name "ML_type") (ml_enclose "val _ = NONE : (" ") option;") #> ml_text (Binding.name "ML_struct") (ml_enclose "functor XXX() = struct structure XX = " " end;") #> diff -r ec2e20b27638 -r 915af80f74b3 src/Pure/thm.ML --- a/src/Pure/thm.ML Sat Feb 11 13:41:36 2012 +0100 +++ b/src/Pure/thm.ML Tue Feb 14 12:40:55 2012 +0100 @@ -1306,8 +1306,8 @@ | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); -(*Increment variables and parameters of orule as required for - resolution with a goal.*) +(*Prepare orule for resolution by lifting it over the parameters and +assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal;