diff -r fb7756087101 -r 38b049cd3aad src/Doc/Implementation/Proof.thy --- a/src/Doc/Implementation/Proof.thy Wed Dec 16 16:31:36 2015 +0100 +++ b/src/Doc/Implementation/Proof.thy Wed Dec 16 17:28:49 2015 +0100 @@ -9,52 +9,48 @@ section \Variables \label{sec:variables}\ text \ - Any variable that is not explicitly bound by \\\-abstraction - is considered as ``free''. Logically, free variables act like - outermost universal quantification at the sequent level: \A\<^sub>1(x), \, A\<^sub>n(x) \ B(x)\ means that the result - holds \<^emph>\for all\ values of \x\. Free variables for - terms (not types) can be fully internalized into the logic: \\ B(x)\ and \\ \x. B(x)\ are interchangeable, provided - that \x\ does not occur elsewhere in the context. - Inspecting \\ \x. B(x)\ more closely, we see that inside the - quantifier, \x\ is essentially ``arbitrary, but fixed'', - while from outside it appears as a place-holder for instantiation - (thanks to \\\ elimination). + Any variable that is not explicitly bound by \\\-abstraction is considered + as ``free''. Logically, free variables act like outermost universal + quantification at the sequent level: \A\<^sub>1(x), \, A\<^sub>n(x) \ B(x)\ means that + the result holds \<^emph>\for all\ values of \x\. Free variables for terms (not + types) can be fully internalized into the logic: \\ B(x)\ and \\ \x. B(x)\ + are interchangeable, provided that \x\ does not occur elsewhere in the + context. Inspecting \\ \x. B(x)\ more closely, we see that inside the + quantifier, \x\ is essentially ``arbitrary, but fixed'', while from outside + it appears as a place-holder for instantiation (thanks to \\\ elimination). - The Pure logic represents the idea of variables being either inside - or outside the current scope by providing separate syntactic - categories for \<^emph>\fixed variables\ (e.g.\ \x\) vs.\ - \<^emph>\schematic variables\ (e.g.\ \?x\). Incidently, a - universal result \\ \x. B(x)\ has the HHF normal form \\ B(?x)\, which represents its generality without requiring an - explicit quantifier. The same principle works for type variables: - \\ B(?\)\ represents the idea of ``\\ \\. B(\)\'' - without demanding a truly polymorphic framework. + The Pure logic represents the idea of variables being either inside or + outside the current scope by providing separate syntactic categories for + \<^emph>\fixed variables\ (e.g.\ \x\) vs.\ \<^emph>\schematic variables\ (e.g.\ \?x\). + Incidently, a universal result \\ \x. B(x)\ has the HHF normal form \\ + B(?x)\, which represents its generality without requiring an explicit + quantifier. The same principle works for type variables: \\ B(?\)\ + represents the idea of ``\\ \\. B(\)\'' without demanding a truly + polymorphic framework. \<^medskip> - Additional care is required to treat type variables in a - way that facilitates type-inference. In principle, term variables - depend on type variables, which means that type variables would have - to be declared first. For example, a raw type-theoretic framework - would demand the context to be constructed in stages as follows: - \\ = \: type, x: \, a: A(x\<^sub>\)\. + Additional care is required to treat type variables in a way that + facilitates type-inference. In principle, term variables depend on type + variables, which means that type variables would have to be declared first. + For example, a raw type-theoretic framework would demand the context to be + constructed in stages as follows: \\ = \: type, x: \, a: A(x\<^sub>\)\. - We allow a slightly less formalistic mode of operation: term - variables \x\ are fixed without specifying a type yet - (essentially \<^emph>\all\ potential occurrences of some instance - \x\<^sub>\\ are fixed); the first occurrence of \x\ + We allow a slightly less formalistic mode of operation: term variables \x\ + are fixed without specifying a type yet (essentially \<^emph>\all\ potential + occurrences of some instance \x\<^sub>\\ are fixed); the first occurrence of \x\ within a specific term assigns its most general type, which is then - maintained consistently in the context. The above example becomes - \\ = x: term, \: type, A(x\<^sub>\)\, where type \\\ is fixed \<^emph>\after\ term \x\, and the constraint - \x :: \\ is an implicit consequence of the occurrence of - \x\<^sub>\\ in the subsequent proposition. + maintained consistently in the context. The above example becomes \\ = x: + term, \: type, A(x\<^sub>\)\, where type \\\ is fixed \<^emph>\after\ term \x\, and the + constraint \x :: \\ is an implicit consequence of the occurrence of \x\<^sub>\\ in + the subsequent proposition. - This twist of dependencies is also accommodated by the reverse - operation of exporting results from a context: a type variable - \\\ is considered fixed as long as it occurs in some fixed - term variable of the context. For example, exporting \x: - term, \: type \ x\<^sub>\ \ x\<^sub>\\ produces in the first step \x: term - \ x\<^sub>\ \ x\<^sub>\\ for fixed \\\, and only in the second step - \\ ?x\<^sub>?\<^sub>\ \ ?x\<^sub>?\<^sub>\\ for schematic \?x\ and \?\\. - The following Isar source text illustrates this scenario. + This twist of dependencies is also accommodated by the reverse operation of + exporting results from a context: a type variable \\\ is considered fixed as + long as it occurs in some fixed term variable of the context. For example, + exporting \x: term, \: type \ x\<^sub>\ \ x\<^sub>\\ produces in the first step \x: term + \ x\<^sub>\ \ x\<^sub>\\ for fixed \\\, and only in the second step \\ ?x\<^sub>?\<^sub>\ \ ?x\<^sub>?\<^sub>\\ + for schematic \?x\ and \?\\. The following Isar source text illustrates this + scenario. \ notepad @@ -70,29 +66,28 @@ thm this \ \fully general result for arbitrary \?x::?'a\\ end -text \The Isabelle/Isar proof context manages the details of term - vs.\ type variables, with high-level principles for moving the - frontier between fixed and schematic variables. +text \ + The Isabelle/Isar proof context manages the details of term vs.\ type + variables, with high-level principles for moving the frontier between fixed + and schematic variables. - The \add_fixes\ operation explicitly declares fixed - variables; the \declare_term\ operation absorbs a term into - a context by fixing new type variables and adding syntactic - constraints. + The \add_fixes\ operation explicitly declares fixed variables; the + \declare_term\ operation absorbs a term into a context by fixing new type + variables and adding syntactic constraints. - The \export\ operation is able to perform the main work of - generalizing term and type variables as sketched above, assuming - that fixing variables and terms have been declared properly. + The \export\ operation is able to perform the main work of generalizing term + and type variables as sketched above, assuming that fixing variables and + terms have been declared properly. - There \import\ operation makes a generalized fact a genuine - part of the context, by inventing fixed variables for the schematic - ones. The effect can be reversed by using \export\ later, - potentially with an extended context; the result is equivalent to - the original modulo renaming of schematic variables. + There \import\ operation makes a generalized fact a genuine part of the + context, by inventing fixed variables for the schematic ones. The effect can + be reversed by using \export\ later, potentially with an extended context; + the result is equivalent to the original modulo renaming of schematic + variables. - The \focus\ operation provides a variant of \import\ - for nested propositions (with explicit quantification): \\x\<^sub>1 \ x\<^sub>n. B(x\<^sub>1, \, x\<^sub>n)\ is - decomposed by inventing fixed variables \x\<^sub>1, \, - x\<^sub>n\ for the body. + The \focus\ operation provides a variant of \import\ for nested propositions + (with explicit quantification): \\x\<^sub>1 \ x\<^sub>n. B(x\<^sub>1, \, x\<^sub>n)\ is decomposed + by inventing fixed variables \x\<^sub>1, \, x\<^sub>n\ for the body. \ text %mlref \ @@ -111,48 +106,47 @@ ((string * (string * typ)) list * term) * Proof.context"} \\ \end{mldecls} - \<^descr> @{ML Variable.add_fixes}~\xs ctxt\ fixes term - variables \xs\, returning the resulting internal names. By - default, the internal representation coincides with the external - one, which also means that the given variables must not be fixed - already. There is a different policy within a local proof body: the - given names are just hints for newly invented Skolem variables. + \<^descr> @{ML Variable.add_fixes}~\xs ctxt\ fixes term variables \xs\, returning + the resulting internal names. By default, the internal representation + coincides with the external one, which also means that the given variables + must not be fixed already. There is a different policy within a local proof + body: the given names are just hints for newly invented Skolem variables. - \<^descr> @{ML Variable.variant_fixes} is similar to @{ML - Variable.add_fixes}, but always produces fresh variants of the given - names. + \<^descr> @{ML Variable.variant_fixes} is similar to @{ML Variable.add_fixes}, but + always produces fresh variants of the given names. - \<^descr> @{ML Variable.declare_term}~\t ctxt\ declares term - \t\ to belong to the context. This automatically fixes new - type variables, but not term variables. Syntactic constraints for - type and term variables are declared uniformly, though. + \<^descr> @{ML Variable.declare_term}~\t ctxt\ declares term \t\ to belong to the + context. This automatically fixes new type variables, but not term + variables. Syntactic constraints for type and term variables are declared + uniformly, though. + + \<^descr> @{ML Variable.declare_constraints}~\t ctxt\ declares syntactic constraints + from term \t\, without making it part of the context yet. - \<^descr> @{ML Variable.declare_constraints}~\t ctxt\ declares - syntactic constraints from term \t\, without making it part - of the context yet. - - \<^descr> @{ML Variable.export}~\inner outer thms\ generalizes - fixed type and term variables in \thms\ according to the - difference of the \inner\ and \outer\ context, - following the principles sketched above. + \<^descr> @{ML Variable.export}~\inner outer thms\ generalizes fixed type and term + variables in \thms\ according to the difference of the \inner\ and \outer\ + context, following the principles sketched above. - \<^descr> @{ML Variable.polymorphic}~\ctxt ts\ generalizes type - variables in \ts\ as far as possible, even those occurring - in fixed term variables. The default policy of type-inference is to - fix newly introduced type variables, which is essentially reversed - with @{ML Variable.polymorphic}: here the given terms are detached - from the context as far as possible. + \<^descr> @{ML Variable.polymorphic}~\ctxt ts\ generalizes type variables in \ts\ as + far as possible, even those occurring in fixed term variables. The default + policy of type-inference is to fix newly introduced type variables, which is + essentially reversed with @{ML Variable.polymorphic}: here the given terms + are detached from the context as far as possible. - \<^descr> @{ML Variable.import}~\open thms ctxt\ invents fixed - type and term variables for the schematic ones occurring in \thms\. The \open\ flag indicates whether the fixed names - should be accessible to the user, otherwise newly introduced names - are marked as ``internal'' (\secref{sec:names}). + \<^descr> @{ML Variable.import}~\open thms ctxt\ invents fixed type and term + variables for the schematic ones occurring in \thms\. The \open\ flag + indicates whether the fixed names should be accessible to the user, + otherwise newly introduced names are marked as ``internal'' + (\secref{sec:names}). - \<^descr> @{ML Variable.focus}~\bindings B\ decomposes the outermost \\\ prefix of proposition \B\, using the given name bindings. + \<^descr> @{ML Variable.focus}~\bindings B\ decomposes the outermost \\\ prefix of + proposition \B\, using the given name bindings. \ -text %mlex \The following example shows how to work with fixed term - and type parameters and with type-inference.\ +text %mlex \ + The following example shows how to work with fixed term and type parameters + and with type-inference. +\ ML_val \ (*static compile-time context -- for testing only*) @@ -173,12 +167,13 @@ val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*) \ -text \In the above example, the starting context is derived from the - toplevel theory, which means that fixed variables are internalized - literally: \x\ is mapped again to \x\, and - attempting to fix it again in the subsequent context is an error. - Alternatively, fixed parameters can be renamed explicitly as - follows:\ +text \ + In the above example, the starting context is derived from the toplevel + theory, which means that fixed variables are internalized literally: \x\ is + mapped again to \x\, and attempting to fix it again in the subsequent + context is an error. Alternatively, fixed parameters can be renamed + explicitly as follows: +\ ML_val \ val ctxt0 = @{context}; @@ -186,11 +181,12 @@ ctxt0 |> Variable.variant_fixes ["x", "x", "x"]; \ -text \The following ML code can now work with the invented names of - \x1\, \x2\, \x3\, without depending on - the details on the system policy for introducing these variants. - Recall that within a proof body the system always invents fresh - ``Skolem constants'', e.g.\ as follows:\ +text \ + The following ML code can now work with the invented names of \x1\, \x2\, + \x3\, without depending on the details on the system policy for introducing + these variants. Recall that within a proof body the system always invents + fresh ``Skolem constants'', e.g.\ as follows: +\ notepad begin @@ -205,68 +201,63 @@ ctxt3 |> Variable.variant_fixes ["y", "y"];\ end -text \In this situation @{ML Variable.add_fixes} and @{ML - Variable.variant_fixes} are very similar, but identical name - proposals given in a row are only accepted by the second version. +text \ + In this situation @{ML Variable.add_fixes} and @{ML Variable.variant_fixes} + are very similar, but identical name proposals given in a row are only + accepted by the second version. \ section \Assumptions \label{sec:assumptions}\ text \ - An \<^emph>\assumption\ is a proposition that it is postulated in the - current context. Local conclusions may use assumptions as - additional facts, but this imposes implicit hypotheses that weaken - the overall statement. + An \<^emph>\assumption\ is a proposition that it is postulated in the current + context. Local conclusions may use assumptions as additional facts, but this + imposes implicit hypotheses that weaken the overall statement. - Assumptions are restricted to fixed non-schematic statements, i.e.\ - all generality needs to be expressed by explicit quantifiers. - Nevertheless, the result will be in HHF normal form with outermost - quantifiers stripped. For example, by assuming \\x :: \. P - x\ we get \\x :: \. P x \ P ?x\ for schematic \?x\ - of fixed type \\\. Local derivations accumulate more and - more explicit references to hypotheses: \A\<^sub>1, \, - A\<^sub>n \ B\ where \A\<^sub>1, \, A\<^sub>n\ needs to - be covered by the assumptions of the current context. + Assumptions are restricted to fixed non-schematic statements, i.e.\ all + generality needs to be expressed by explicit quantifiers. Nevertheless, the + result will be in HHF normal form with outermost quantifiers stripped. For + example, by assuming \\x :: \. P x\ we get \\x :: \. P x \ P ?x\ for + schematic \?x\ of fixed type \\\. Local derivations accumulate more and more + explicit references to hypotheses: \A\<^sub>1, \, A\<^sub>n \ B\ where \A\<^sub>1, \, A\<^sub>n\ + needs to be covered by the assumptions of the current context. \<^medskip> - The \add_assms\ operation augments the context by - local assumptions, which are parameterized by an arbitrary \export\ rule (see below). + The \add_assms\ operation augments the context by local assumptions, which + are parameterized by an arbitrary \export\ rule (see below). - The \export\ operation moves facts from a (larger) inner - context into a (smaller) outer context, by discharging the - difference of the assumptions as specified by the associated export - rules. Note that the discharged portion is determined by the - difference of contexts, not the facts being exported! There is a - separate flag to indicate a goal context, where the result is meant - to refine an enclosing sub-goal of a structured proof state. + The \export\ operation moves facts from a (larger) inner context into a + (smaller) outer context, by discharging the difference of the assumptions as + specified by the associated export rules. Note that the discharged portion + is determined by the difference of contexts, not the facts being exported! + There is a separate flag to indicate a goal context, where the result is + meant to refine an enclosing sub-goal of a structured proof state. \<^medskip> - The most basic export rule discharges assumptions directly - by means of the \\\ introduction rule: + The most basic export rule discharges assumptions directly by means of the + \\\ introduction rule: \[ \infer[(\\\intro\)]{\\ - A \ A \ B\}{\\ \ B\} \] - The variant for goal refinements marks the newly introduced - premises, which causes the canonical Isar goal refinement scheme to - enforce unification with local premises within the goal: + The variant for goal refinements marks the newly introduced premises, which + causes the canonical Isar goal refinement scheme to enforce unification with + local premises within the goal: \[ \infer[(\#\\intro\)]{\\ - A \ #A \ B\}{\\ \ B\} \] \<^medskip> - Alternative versions of assumptions may perform arbitrary - transformations on export, as long as the corresponding portion of - hypotheses is removed from the given facts. For example, a local - definition works by fixing \x\ and assuming \x \ t\, - with the following export rule to reverse the effect: + Alternative versions of assumptions may perform arbitrary transformations on + export, as long as the corresponding portion of hypotheses is removed from + the given facts. For example, a local definition works by fixing \x\ and + assuming \x \ t\, with the following export rule to reverse the effect: \[ \infer[(\\\expand\)]{\\ - (x \ t) \ B t\}{\\ \ B x\} \] - This works, because the assumption \x \ t\ was introduced in - a context with \x\ being fresh, so \x\ does not - occur in \\\ here. + This works, because the assumption \x \ t\ was introduced in a context with + \x\ being fresh, so \x\ does not occur in \\\ here. \ text %mlref \ @@ -281,36 +272,33 @@ @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\ \end{mldecls} - \<^descr> Type @{ML_type Assumption.export} represents arbitrary export - rules, which is any function of type @{ML_type "bool -> cterm list - -> thm -> thm"}, where the @{ML_type "bool"} indicates goal mode, - and the @{ML_type "cterm list"} the collection of assumptions to be - discharged simultaneously. + \<^descr> Type @{ML_type Assumption.export} represents arbitrary export rules, which + is any function of type @{ML_type "bool -> cterm list -> thm -> thm"}, where + the @{ML_type "bool"} indicates goal mode, and the @{ML_type "cterm list"} + the collection of assumptions to be discharged simultaneously. - \<^descr> @{ML Assumption.assume}~\ctxt A\ turns proposition \A\ into a primitive assumption \A \ A'\, where the - conclusion \A'\ is in HHF normal form. + \<^descr> @{ML Assumption.assume}~\ctxt A\ turns proposition \A\ into a primitive + assumption \A \ A'\, where the conclusion \A'\ is in HHF normal form. - \<^descr> @{ML Assumption.add_assms}~\r As\ augments the context - by assumptions \As\ with export rule \r\. The - resulting facts are hypothetical theorems as produced by the raw - @{ML Assumption.assume}. + \<^descr> @{ML Assumption.add_assms}~\r As\ augments the context by assumptions \As\ + with export rule \r\. The resulting facts are hypothetical theorems as + produced by the raw @{ML Assumption.assume}. - \<^descr> @{ML Assumption.add_assumes}~\As\ is a special case of - @{ML Assumption.add_assms} where the export rule performs \\\intro\ or \#\\intro\, depending on goal - mode. + \<^descr> @{ML Assumption.add_assumes}~\As\ is a special case of @{ML + Assumption.add_assms} where the export rule performs \\\intro\ or + \#\\intro\, depending on goal mode. - \<^descr> @{ML Assumption.export}~\is_goal inner outer thm\ - exports result \thm\ from the the \inner\ context - back into the \outer\ one; \is_goal = true\ means - this is a goal context. The result is in HHF normal form. Note - that @{ML "Proof_Context.export"} combines @{ML "Variable.export"} - and @{ML "Assumption.export"} in the canonical way. + \<^descr> @{ML Assumption.export}~\is_goal inner outer thm\ exports result \thm\ + from the the \inner\ context back into the \outer\ one; \is_goal = true\ + means this is a goal context. The result is in HHF normal form. Note that + @{ML "Proof_Context.export"} combines @{ML "Variable.export"} and @{ML + "Assumption.export"} in the canonical way. \ -text %mlex \The following example demonstrates how rules can be - derived by building up a context of assumptions first, and exporting - some local fact afterwards. We refer to @{theory Pure} equality - here for testing purposes. +text %mlex \ + The following example demonstrates how rules can be derived by building up a + context of assumptions first, and exporting some local fact afterwards. We + refer to @{theory Pure} equality here for testing purposes. \ ML_val \ @@ -325,52 +313,50 @@ val r = Assumption.export false ctxt1 ctxt0 eq'; \ -text \Note that the variables of the resulting rule are not - generalized. This would have required to fix them properly in the - context beforehand, and export wrt.\ variables afterwards (cf.\ @{ML - Variable.export} or the combined @{ML "Proof_Context.export"}).\ +text \ + Note that the variables of the resulting rule are not generalized. This + would have required to fix them properly in the context beforehand, and + export wrt.\ variables afterwards (cf.\ @{ML Variable.export} or the + combined @{ML "Proof_Context.export"}). +\ section \Structured goals and results \label{sec:struct-goals}\ text \ - Local results are established by monotonic reasoning from facts - within a context. This allows common combinations of theorems, - e.g.\ via \\/\\ elimination, resolution rules, or equational - reasoning, see \secref{sec:thms}. Unaccounted context manipulations - should be avoided, notably raw \\/\\ introduction or ad-hoc - references to free variables or assumptions not present in the proof - context. + Local results are established by monotonic reasoning from facts within a + context. This allows common combinations of theorems, e.g.\ via \\/\\ + elimination, resolution rules, or equational reasoning, see + \secref{sec:thms}. Unaccounted context manipulations should be avoided, + notably raw \\/\\ introduction or ad-hoc references to free variables or + assumptions not present in the proof context. \<^medskip> - The \SUBPROOF\ combinator allows to structure a - tactical proof recursively by decomposing a selected sub-goal: - \(\x. A(x) \ B(x)) \ \\ is turned into \B(x) \ \\ - after fixing \x\ and assuming \A(x)\. This means - the tactic needs to solve the conclusion, but may use the premise as - a local fact, for locally fixed variables. + The \SUBPROOF\ combinator allows to structure a tactical proof recursively + by decomposing a selected sub-goal: \(\x. A(x) \ B(x)) \ \\ is turned into + \B(x) \ \\ after fixing \x\ and assuming \A(x)\. This means the tactic needs + to solve the conclusion, but may use the premise as a local fact, for + locally fixed variables. - The family of \FOCUS\ combinators is similar to \SUBPROOF\, but allows to retain schematic variables and pending - subgoals in the resulting goal state. + The family of \FOCUS\ combinators is similar to \SUBPROOF\, but allows to + retain schematic variables and pending subgoals in the resulting goal state. - The \prove\ operation provides an interface for structured - backwards reasoning under program control, with some explicit sanity - checks of the result. The goal context can be augmented by - additional fixed variables (cf.\ \secref{sec:variables}) and - assumptions (cf.\ \secref{sec:assumptions}), which will be available - as local facts during the proof and discharged into implications in - the result. Type and term variables are generalized as usual, - according to the context. + The \prove\ operation provides an interface for structured backwards + reasoning under program control, with some explicit sanity checks of the + result. The goal context can be augmented by additional fixed variables + (cf.\ \secref{sec:variables}) and assumptions (cf.\ + \secref{sec:assumptions}), which will be available as local facts during the + proof and discharged into implications in the result. Type and term + variables are generalized as usual, according to the context. - The \obtain\ operation produces results by eliminating - existing facts by means of a given tactic. This acts like a dual - conclusion: the proof demonstrates that the context may be augmented - by parameters and assumptions, without affecting any conclusions - that do not mention these parameters. See also - @{cite "isabelle-isar-ref"} for the user-level @{command obtain} and - @{command guess} elements. Final results, which may not refer to - the parameters in the conclusion, need to exported explicitly into - the original context.\ + The \obtain\ operation produces results by eliminating existing facts by + means of a given tactic. This acts like a dual conclusion: the proof + demonstrates that the context may be augmented by parameters and + assumptions, without affecting any conclusions that do not mention these + parameters. See also @{cite "isabelle-isar-ref"} for the user-level + @{command obtain} and @{command guess} elements. Final results, which may + not refer to the parameters in the conclusion, need to exported explicitly + into the original context.\ text %mlref \ \begin{mldecls} @@ -402,54 +388,54 @@ Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ \end{mldecls} - \<^descr> @{ML SUBPROOF}~\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 - schematic parameters of the goal are imported into the context as - fixed ones, which may not be instantiated in the sub-proof. + \<^descr> @{ML SUBPROOF}~\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 schematic parameters of the goal are + imported into the context as fixed ones, which may not be instantiated in + the sub-proof. \<^descr> @{ML Subgoal.FOCUS}, @{ML Subgoal.FOCUS_PREMS}, and @{ML - Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are - slightly more flexible: only the specified parts of the subgoal are - imported into the context, and the body tactic may introduce new - subgoals and schematic variables. + Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are slightly more + flexible: only the specified parts of the subgoal are imported into the + context, and the body tactic may introduce new subgoals and schematic + variables. - \<^descr> @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML - Subgoal.focus_params} extract the focus information from a goal - state in the same way as the corresponding tacticals above. This is - occasionally useful to experiment without writing actual tactics - yet. + \<^descr> @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML Subgoal.focus_params} + extract the focus information from a goal state in the same way as the + corresponding tacticals above. This is occasionally useful to experiment + without writing actual tactics yet. - \<^descr> @{ML Goal.prove}~\ctxt xs As C tac\ states goal \C\ in the context augmented by fixed variables \xs\ and - assumptions \As\, and applies tactic \tac\ to solve - it. The latter may depend on the local assumptions being presented - as facts. The result is in HHF normal form. + \<^descr> @{ML Goal.prove}~\ctxt xs As C tac\ states goal \C\ in the context + augmented by fixed variables \xs\ and assumptions \As\, and applies tactic + \tac\ to solve it. The latter may depend on the local assumptions being + presented as facts. The result is in HHF normal form. - \<^descr> @{ML Goal.prove_common}~\ctxt fork_pri\ is the common form - to state and prove a simultaneous goal statement, where @{ML Goal.prove} - is a convenient shorthand that is most frequently used in applications. + \<^descr> @{ML Goal.prove_common}~\ctxt fork_pri\ is the common form to state and + prove a simultaneous goal statement, where @{ML Goal.prove} is a convenient + shorthand that is most frequently used in applications. The given list of simultaneous conclusions is encoded in the goal state by - means of Pure conjunction: @{ML Goal.conjunction_tac} will turn this into - a collection of individual subgoals, but note that the original multi-goal + means of Pure conjunction: @{ML Goal.conjunction_tac} will turn this into a + collection of individual subgoals, but note that the original multi-goal state is usually required for advanced induction. - It is possible to provide an optional priority for a forked proof, - typically @{ML "SOME ~1"}, while @{ML NONE} means the proof is immediate - (sequential) as for @{ML Goal.prove}. Note that a forked proof does not - exhibit any failures in the usual way via exceptions in ML, but - accumulates error situations under the execution id of the running - transaction. Thus the system is able to expose error messages ultimately - to the end-user, even though the subsequent ML code misses them. + It is possible to provide an optional priority for a forked proof, typically + @{ML "SOME ~1"}, while @{ML NONE} means the proof is immediate (sequential) + as for @{ML Goal.prove}. Note that a forked proof does not exhibit any + failures in the usual way via exceptions in ML, but accumulates error + situations under the execution id of the running transaction. Thus the + system is able to expose error messages ultimately to the end-user, even + though the subsequent ML code misses them. - \<^descr> @{ML Obtain.result}~\tac thms ctxt\ eliminates the - given facts using a tactic, which results in additional fixed - variables and assumptions in the context. Final results need to be - exported explicitly. + \<^descr> @{ML Obtain.result}~\tac thms ctxt\ eliminates the given facts using a + tactic, which results in additional fixed variables and assumptions in the + context. Final results need to be exported explicitly. \ -text %mlex \The following minimal example illustrates how to access - the focus information of a structured goal state.\ +text %mlex \ + The following minimal example illustrates how to access the focus + information of a structured goal state. +\ notepad begin @@ -467,8 +453,9 @@ text \ \<^medskip> - The next example demonstrates forward-elimination in - a local context, using @{ML Obtain.result}.\ + The next example demonstrates forward-elimination in a local context, using + @{ML Obtain.result}. +\ notepad begin