| author | wenzelm | 
| Thu, 05 Mar 2009 02:20:06 +0100 | |
| changeset 30270 | 61811c9224a6 | 
| parent 29761 | 2b658e50683a | 
| child 30272 | 2d612824e642 | 
| permissions | -rw-r--r-- | 
| 29755 | 1  | 
theory Proof  | 
2  | 
imports Base  | 
|
3  | 
begin  | 
|
| 18537 | 4  | 
|
| 20451 | 5  | 
chapter {* Structured proofs *}
 | 
| 18537 | 6  | 
|
| 20474 | 7  | 
section {* Variables \label{sec:variables} *}
 | 
| 20026 | 8  | 
|
| 20470 | 9  | 
text {*
 | 
10  | 
  Any variable that is not explicitly bound by @{text "\<lambda>"}-abstraction
 | 
|
11  | 
is considered as ``free''. Logically, free variables act like  | 
|
| 20474 | 12  | 
  outermost universal quantification at the sequent level: @{text
 | 
| 20470 | 13  | 
"A\<^isub>1(x), \<dots>, A\<^isub>n(x) \<turnstile> B(x)"} means that the result  | 
14  | 
  holds \emph{for all} values of @{text "x"}.  Free variables for
 | 
|
15  | 
  terms (not types) can be fully internalized into the logic: @{text
 | 
|
| 20474 | 16  | 
  "\<turnstile> B(x)"} and @{text "\<turnstile> \<And>x. B(x)"} are interchangeable, provided
 | 
17  | 
  that @{text "x"} does not occur elsewhere in the context.
 | 
|
18  | 
  Inspecting @{text "\<turnstile> \<And>x. B(x)"} more closely, we see that inside the
 | 
|
| 20470 | 19  | 
  quantifier, @{text "x"} is essentially ``arbitrary, but fixed'',
 | 
20  | 
while from outside it appears as a place-holder for instantiation  | 
|
| 20474 | 21  | 
  (thanks to @{text "\<And>"} elimination).
 | 
| 20470 | 22  | 
|
| 20474 | 23  | 
The Pure logic represents the idea of variables being either inside  | 
24  | 
or outside the current scope by providing separate syntactic  | 
|
| 20470 | 25  | 
  categories for \emph{fixed variables} (e.g.\ @{text "x"}) vs.\
 | 
26  | 
  \emph{schematic variables} (e.g.\ @{text "?x"}).  Incidently, a
 | 
|
| 20474 | 27  | 
  universal result @{text "\<turnstile> \<And>x. B(x)"} has the HHF normal form @{text
 | 
| 20470 | 28  | 
"\<turnstile> B(?x)"}, which represents its generality nicely without requiring  | 
| 20474 | 29  | 
an explicit quantifier. The same principle works for type  | 
30  | 
  variables: @{text "\<turnstile> B(?\<alpha>)"} represents the idea of ``@{text "\<turnstile>
 | 
|
| 20470 | 31  | 
\<forall>\<alpha>. B(\<alpha>)"}'' without demanding a truly polymorphic framework.  | 
32  | 
||
33  | 
\medskip Additional care is required to treat type variables in a  | 
|
34  | 
way that facilitates type-inference. In principle, term variables  | 
|
35  | 
depend on type variables, which means that type variables would have  | 
|
36  | 
to be declared first. For example, a raw type-theoretic framework  | 
|
37  | 
would demand the context to be constructed in stages as follows:  | 
|
38  | 
  @{text "\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^isub>\<alpha>)"}.
 | 
|
39  | 
||
40  | 
We allow a slightly less formalistic mode of operation: term  | 
|
41  | 
  variables @{text "x"} are fixed without specifying a type yet
 | 
|
42  | 
  (essentially \emph{all} potential occurrences of some instance
 | 
|
| 20474 | 43  | 
  @{text "x\<^isub>\<tau>"} are fixed); the first occurrence of @{text "x"}
 | 
44  | 
within a specific term assigns its most general type, which is then  | 
|
45  | 
maintained consistently in the context. The above example becomes  | 
|
46  | 
  @{text "\<Gamma> = x: term, \<alpha>: type, A(x\<^isub>\<alpha>)"}, where type @{text
 | 
|
47  | 
  "\<alpha>"} is fixed \emph{after} term @{text "x"}, and the constraint
 | 
|
48  | 
  @{text "x :: \<alpha>"} is an implicit consequence of the occurrence of
 | 
|
49  | 
  @{text "x\<^isub>\<alpha>"} in the subsequent proposition.
 | 
|
| 20470 | 50  | 
|
51  | 
This twist of dependencies is also accommodated by the reverse  | 
|
52  | 
operation of exporting results from a context: a type variable  | 
|
53  | 
  @{text "\<alpha>"} is considered fixed as long as it occurs in some fixed
 | 
|
54  | 
  term variable of the context.  For example, exporting @{text "x:
 | 
|
| 20474 | 55  | 
term, \<alpha>: type \<turnstile> x\<^isub>\<alpha> = x\<^isub>\<alpha>"} produces in the first step  | 
56  | 
  @{text "x: term \<turnstile> x\<^isub>\<alpha> = x\<^isub>\<alpha>"} for fixed @{text "\<alpha>"},
 | 
|
57  | 
  and only in the second step @{text "\<turnstile> ?x\<^isub>?\<^isub>\<alpha> =
 | 
|
58  | 
  ?x\<^isub>?\<^isub>\<alpha>"} for schematic @{text "?x"} and @{text "?\<alpha>"}.
 | 
|
| 20470 | 59  | 
|
60  | 
\medskip The Isabelle/Isar proof context manages the gory details of  | 
|
61  | 
term vs.\ type variables, with high-level principles for moving the  | 
|
| 20474 | 62  | 
frontier between fixed and schematic variables.  | 
63  | 
||
64  | 
  The @{text "add_fixes"} operation explictly declares fixed
 | 
|
65  | 
  variables; the @{text "declare_term"} operation absorbs a term into
 | 
|
66  | 
a context by fixing new type variables and adding syntactic  | 
|
67  | 
constraints.  | 
|
| 20470 | 68  | 
|
| 20474 | 69  | 
  The @{text "export"} operation is able to perform the main work of
 | 
70  | 
generalizing term and type variables as sketched above, assuming  | 
|
71  | 
that fixing variables and terms have been declared properly.  | 
|
72  | 
||
73  | 
  There @{text "import"} operation makes a generalized fact a genuine
 | 
|
74  | 
part of the context, by inventing fixed variables for the schematic  | 
|
75  | 
  ones.  The effect can be reversed by using @{text "export"} later,
 | 
|
76  | 
potentially with an extended context; the result is equivalent to  | 
|
77  | 
the original modulo renaming of schematic variables.  | 
|
| 20470 | 78  | 
|
79  | 
  The @{text "focus"} operation provides a variant of @{text "import"}
 | 
|
80  | 
  for nested propositions (with explicit quantification): @{text
 | 
|
| 20474 | 81  | 
"\<And>x\<^isub>1 \<dots> x\<^isub>n. B(x\<^isub>1, \<dots>, x\<^isub>n)"} is  | 
82  | 
  decomposed by inventing fixed variables @{text "x\<^isub>1, \<dots>,
 | 
|
83  | 
x\<^isub>n"} for the body.  | 
|
| 20470 | 84  | 
*}  | 
| 20064 | 85  | 
|
| 20026 | 86  | 
text %mlref {*
 | 
87  | 
  \begin{mldecls}
 | 
|
| 20474 | 88  | 
  @{index_ML Variable.add_fixes: "
 | 
89  | 
string list -> Proof.context -> string list * Proof.context"} \\  | 
|
| 
20797
 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 
wenzelm 
parents: 
20547 
diff
changeset
 | 
90  | 
  @{index_ML Variable.variant_fixes: "
 | 
| 20474 | 91  | 
string list -> Proof.context -> string list * Proof.context"} \\  | 
| 20026 | 92  | 
  @{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\
 | 
| 20470 | 93  | 
  @{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\
 | 
94  | 
  @{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
 | 
|
95  | 
  @{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
 | 
|
| 
22568
 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 
wenzelm 
parents: 
21827 
diff
changeset
 | 
96  | 
  @{index_ML Variable.import_thms: "bool -> thm list -> Proof.context ->
 | 
| 20474 | 97  | 
((ctyp list * cterm list) * thm list) * Proof.context"} \\  | 
| 20470 | 98  | 
  @{index_ML Variable.focus: "cterm -> Proof.context -> (cterm list * cterm) * Proof.context"} \\
 | 
| 20026 | 99  | 
  \end{mldecls}
 | 
100  | 
||
101  | 
  \begin{description}
 | 
|
102  | 
||
| 20064 | 103  | 
  \item @{ML Variable.add_fixes}~@{text "xs ctxt"} fixes term
 | 
| 20470 | 104  | 
  variables @{text "xs"}, returning the resulting internal names.  By
 | 
105  | 
default, the internal representation coincides with the external  | 
|
| 20474 | 106  | 
one, which also means that the given variables must not be fixed  | 
107  | 
already. There is a different policy within a local proof body: the  | 
|
108  | 
given names are just hints for newly invented Skolem variables.  | 
|
| 20064 | 109  | 
|
| 
20797
 
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
 
wenzelm 
parents: 
20547 
diff
changeset
 | 
110  | 
  \item @{ML Variable.variant_fixes} is similar to @{ML
 | 
| 20470 | 111  | 
Variable.add_fixes}, but always produces fresh variants of the given  | 
| 20474 | 112  | 
names.  | 
| 20470 | 113  | 
|
114  | 
  \item @{ML Variable.declare_term}~@{text "t ctxt"} declares term
 | 
|
115  | 
  @{text "t"} to belong to the context.  This automatically fixes new
 | 
|
116  | 
type variables, but not term variables. Syntactic constraints for  | 
|
| 20474 | 117  | 
type and term variables are declared uniformly, though.  | 
| 20470 | 118  | 
|
| 20474 | 119  | 
  \item @{ML Variable.declare_constraints}~@{text "t ctxt"} declares
 | 
120  | 
  syntactic constraints from term @{text "t"}, without making it part
 | 
|
121  | 
of the context yet.  | 
|
| 20026 | 122  | 
|
| 20470 | 123  | 
  \item @{ML Variable.export}~@{text "inner outer thms"} generalizes
 | 
124  | 
  fixed type and term variables in @{text "thms"} according to the
 | 
|
125  | 
  difference of the @{text "inner"} and @{text "outer"} context,
 | 
|
126  | 
following the principles sketched above.  | 
|
127  | 
||
128  | 
  \item @{ML Variable.polymorphic}~@{text "ctxt ts"} generalizes type
 | 
|
129  | 
  variables in @{text "ts"} as far as possible, even those occurring
 | 
|
130  | 
in fixed term variables. The default policy of type-inference is to  | 
|
| 20474 | 131  | 
fix newly introduced type variables, which is essentially reversed  | 
132  | 
  with @{ML Variable.polymorphic}: here the given terms are detached
 | 
|
133  | 
from the context as far as possible.  | 
|
| 20470 | 134  | 
|
| 
22568
 
ed7aa5a350ef
renamed Variable.import to import_thms (avoid clash with Alice keywords);
 
wenzelm 
parents: 
21827 
diff
changeset
 | 
135  | 
  \item @{ML Variable.import_thms}~@{text "open thms ctxt"} invents fixed
 | 
| 20474 | 136  | 
  type and term variables for the schematic ones occurring in @{text
 | 
137  | 
  "thms"}.  The @{text "open"} flag indicates whether the fixed names
 | 
|
138  | 
should be accessible to the user, otherwise newly introduced names  | 
|
139  | 
  are marked as ``internal'' (\secref{sec:names}).
 | 
|
| 20026 | 140  | 
|
| 20474 | 141  | 
  \item @{ML Variable.focus}~@{text B} decomposes the outermost @{text
 | 
142  | 
  "\<And>"} prefix of proposition @{text "B"}.
 | 
|
| 20026 | 143  | 
|
144  | 
  \end{description}
 | 
|
145  | 
*}  | 
|
146  | 
||
| 18537 | 147  | 
|
| 20474 | 148  | 
section {* Assumptions \label{sec:assumptions} *}
 | 
| 20451 | 149  | 
|
| 20458 | 150  | 
text {*
 | 
151  | 
  An \emph{assumption} is a proposition that it is postulated in the
 | 
|
152  | 
current context. Local conclusions may use assumptions as  | 
|
153  | 
additional facts, but this imposes implicit hypotheses that weaken  | 
|
154  | 
the overall statement.  | 
|
155  | 
||
| 20474 | 156  | 
Assumptions are restricted to fixed non-schematic statements, i.e.\  | 
157  | 
all generality needs to be expressed by explicit quantifiers.  | 
|
| 20458 | 158  | 
Nevertheless, the result will be in HHF normal form with outermost  | 
159  | 
  quantifiers stripped.  For example, by assuming @{text "\<And>x :: \<alpha>. P
 | 
|
| 20474 | 160  | 
  x"} we get @{text "\<And>x :: \<alpha>. P x \<turnstile> P ?x"} for schematic @{text "?x"}
 | 
161  | 
  of fixed type @{text "\<alpha>"}.  Local derivations accumulate more and
 | 
|
162  | 
  more explicit references to hypotheses: @{text "A\<^isub>1, \<dots>,
 | 
|
| 20458 | 163  | 
  A\<^isub>n \<turnstile> B"} where @{text "A\<^isub>1, \<dots>, A\<^isub>n"} needs to
 | 
164  | 
be covered by the assumptions of the current context.  | 
|
165  | 
||
| 20459 | 166  | 
  \medskip The @{text "add_assms"} operation augments the context by
 | 
167  | 
  local assumptions, which are parameterized by an arbitrary @{text
 | 
|
168  | 
"export"} rule (see below).  | 
|
| 20458 | 169  | 
|
170  | 
  The @{text "export"} operation moves facts from a (larger) inner
 | 
|
171  | 
context into a (smaller) outer context, by discharging the  | 
|
172  | 
difference of the assumptions as specified by the associated export  | 
|
173  | 
rules. Note that the discharged portion is determined by the  | 
|
| 20459 | 174  | 
difference contexts, not the facts being exported! There is a  | 
175  | 
separate flag to indicate a goal context, where the result is meant  | 
|
| 29760 | 176  | 
to refine an enclosing sub-goal of a structured proof state.  | 
| 20458 | 177  | 
|
178  | 
\medskip The most basic export rule discharges assumptions directly  | 
|
179  | 
  by means of the @{text "\<Longrightarrow>"} introduction rule:
 | 
|
180  | 
\[  | 
|
181  | 
  \infer[(@{text "\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
 | 
|
182  | 
\]  | 
|
183  | 
||
184  | 
The variant for goal refinements marks the newly introduced  | 
|
| 20474 | 185  | 
premises, which causes the canonical Isar goal refinement scheme to  | 
| 20458 | 186  | 
enforce unification with local premises within the goal:  | 
187  | 
\[  | 
|
188  | 
  \infer[(@{text "#\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
 | 
|
189  | 
\]  | 
|
190  | 
||
| 20474 | 191  | 
\medskip Alternative versions of assumptions may perform arbitrary  | 
192  | 
transformations on export, as long as the corresponding portion of  | 
|
| 20459 | 193  | 
hypotheses is removed from the given facts. For example, a local  | 
194  | 
  definition works by fixing @{text "x"} and assuming @{text "x \<equiv> t"},
 | 
|
195  | 
with the following export rule to reverse the effect:  | 
|
| 20458 | 196  | 
\[  | 
| 20491 | 197  | 
  \infer[(@{text "\<equiv>-expand"})]{@{text "\<Gamma> \\ x \<equiv> t \<turnstile> B t"}}{@{text "\<Gamma> \<turnstile> B x"}}
 | 
| 20458 | 198  | 
\]  | 
| 20474 | 199  | 
  This works, because the assumption @{text "x \<equiv> t"} was introduced in
 | 
200  | 
  a context with @{text "x"} being fresh, so @{text "x"} does not
 | 
|
201  | 
  occur in @{text "\<Gamma>"} here.
 | 
|
| 20458 | 202  | 
*}  | 
203  | 
||
204  | 
text %mlref {*
 | 
|
205  | 
  \begin{mldecls}
 | 
|
206  | 
  @{index_ML_type Assumption.export} \\
 | 
|
207  | 
  @{index_ML Assumption.assume: "cterm -> thm"} \\
 | 
|
208  | 
  @{index_ML Assumption.add_assms:
 | 
|
| 20459 | 209  | 
"Assumption.export ->  | 
210  | 
cterm list -> Proof.context -> thm list * Proof.context"} \\  | 
|
211  | 
  @{index_ML Assumption.add_assumes: "
 | 
|
212  | 
cterm list -> Proof.context -> thm list * Proof.context"} \\  | 
|
| 20458 | 213  | 
  @{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
 | 
214  | 
  \end{mldecls}
 | 
|
215  | 
||
216  | 
  \begin{description}
 | 
|
217  | 
||
| 20459 | 218  | 
  \item @{ML_type Assumption.export} represents arbitrary export
 | 
219  | 
  rules, which is any function of type @{ML_type "bool -> cterm list -> thm -> thm"},
 | 
|
220  | 
  where the @{ML_type "bool"} indicates goal mode, and the @{ML_type
 | 
|
221  | 
"cterm list"} the collection of assumptions to be discharged  | 
|
222  | 
simultaneously.  | 
|
| 20458 | 223  | 
|
| 20459 | 224  | 
  \item @{ML Assumption.assume}~@{text "A"} turns proposition @{text
 | 
225  | 
  "A"} into a raw assumption @{text "A \<turnstile> A'"}, where the conclusion
 | 
|
226  | 
  @{text "A'"} is in HHF normal form.
 | 
|
| 20458 | 227  | 
|
| 20474 | 228  | 
  \item @{ML Assumption.add_assms}~@{text "r As"} augments the context
 | 
229  | 
  by assumptions @{text "As"} with export rule @{text "r"}.  The
 | 
|
230  | 
resulting facts are hypothetical theorems as produced by the raw  | 
|
231  | 
  @{ML Assumption.assume}.
 | 
|
| 20459 | 232  | 
|
233  | 
  \item @{ML Assumption.add_assumes}~@{text "As"} is a special case of
 | 
|
234  | 
  @{ML Assumption.add_assms} where the export rule performs @{text
 | 
|
235  | 
  "\<Longrightarrow>_intro"} or @{text "#\<Longrightarrow>_intro"}, depending on goal mode.
 | 
|
| 20458 | 236  | 
|
| 20474 | 237  | 
  \item @{ML Assumption.export}~@{text "is_goal inner outer thm"}
 | 
238  | 
  exports result @{text "thm"} from the the @{text "inner"} context
 | 
|
| 20459 | 239  | 
  back into the @{text "outer"} one; @{text "is_goal = true"} means
 | 
240  | 
this is a goal context. The result is in HHF normal form. Note  | 
|
241  | 
  that @{ML "ProofContext.export"} combines @{ML "Variable.export"}
 | 
|
242  | 
  and @{ML "Assumption.export"} in the canonical way.
 | 
|
| 20458 | 243  | 
|
244  | 
  \end{description}
 | 
|
245  | 
*}  | 
|
| 20451 | 246  | 
|
247  | 
||
| 28541 | 248  | 
section {* Results \label{sec:results} *}
 | 
| 20451 | 249  | 
|
| 18537 | 250  | 
text {*
 | 
| 20472 | 251  | 
Local results are established by monotonic reasoning from facts  | 
| 20474 | 252  | 
within a context. This allows common combinations of theorems,  | 
253  | 
  e.g.\ via @{text "\<And>/\<Longrightarrow>"} elimination, resolution rules, or equational
 | 
|
254  | 
  reasoning, see \secref{sec:thms}.  Unaccounted context manipulations
 | 
|
255  | 
  should be avoided, notably raw @{text "\<And>/\<Longrightarrow>"} introduction or ad-hoc
 | 
|
| 20472 | 256  | 
references to free variables or assumptions not present in the proof  | 
257  | 
context.  | 
|
| 18537 | 258  | 
|
| 20491 | 259  | 
  \medskip The @{text "SUBPROOF"} combinator allows to structure a
 | 
260  | 
tactical proof recursively by decomposing a selected sub-goal:  | 
|
261  | 
  @{text "(\<And>x. A(x) \<Longrightarrow> B(x)) \<Longrightarrow> \<dots>"} is turned into @{text "B(x) \<Longrightarrow> \<dots>"}
 | 
|
262  | 
  after fixing @{text "x"} and assuming @{text "A(x)"}.  This means
 | 
|
263  | 
the tactic needs to solve the conclusion, but may use the premise as  | 
|
264  | 
a local fact, for locally fixed variables.  | 
|
| 18537 | 265  | 
|
| 20491 | 266  | 
  The @{text "prove"} operation provides an interface for structured
 | 
267  | 
backwards reasoning under program control, with some explicit sanity  | 
|
268  | 
checks of the result. The goal context can be augmented by  | 
|
269  | 
  additional fixed variables (cf.\ \secref{sec:variables}) and
 | 
|
270  | 
  assumptions (cf.\ \secref{sec:assumptions}), which will be available
 | 
|
271  | 
as local facts during the proof and discharged into implications in  | 
|
272  | 
the result. Type and term variables are generalized as usual,  | 
|
273  | 
according to the context.  | 
|
| 18537 | 274  | 
|
| 20491 | 275  | 
  The @{text "obtain"} operation produces results by eliminating
 | 
276  | 
existing facts by means of a given tactic. This acts like a dual  | 
|
277  | 
conclusion: the proof demonstrates that the context may be augmented  | 
|
278  | 
by certain fixed variables and assumptions. See also  | 
|
279  | 
  \cite{isabelle-isar-ref} for the user-level @{text "\<OBTAIN>"} and
 | 
|
280  | 
  @{text "\<GUESS>"} elements.  Final results, which may not refer to
 | 
|
281  | 
the parameters in the conclusion, need to exported explicitly into  | 
|
282  | 
the original context.  | 
|
| 18537 | 283  | 
*}  | 
284  | 
||
| 20472 | 285  | 
text %mlref {*
 | 
286  | 
  \begin{mldecls}
 | 
|
| 20491 | 287  | 
  @{index_ML SUBPROOF:
 | 
288  | 
  "({context: Proof.context, schematics: ctyp list * cterm list,
 | 
|
289  | 
params: cterm list, asms: cterm list, concl: cterm,  | 
|
290  | 
prems: thm list} -> tactic) -> Proof.context -> int -> tactic"} \\  | 
|
| 20547 | 291  | 
  \end{mldecls}
 | 
292  | 
  \begin{mldecls}
 | 
|
| 20472 | 293  | 
  @{index_ML Goal.prove: "Proof.context -> string list -> term list -> term ->
 | 
294  | 
  ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\
 | 
|
295  | 
  @{index_ML Goal.prove_multi: "Proof.context -> string list -> term list -> term list ->
 | 
|
296  | 
  ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\
 | 
|
| 20547 | 297  | 
  \end{mldecls}
 | 
298  | 
  \begin{mldecls}
 | 
|
| 20491 | 299  | 
  @{index_ML Obtain.result: "(Proof.context -> tactic) ->
 | 
| 20542 | 300  | 
thm list -> Proof.context -> (cterm list * thm list) * Proof.context"} \\  | 
| 20472 | 301  | 
  \end{mldecls}
 | 
| 18537 | 302  | 
|
| 20472 | 303  | 
  \begin{description}
 | 
| 18537 | 304  | 
|
| 29761 | 305  | 
  \item @{ML SUBPROOF}~@{text "tac ctxt i"} decomposes the structure
 | 
306  | 
of the specified sub-goal, producing an extended context and a  | 
|
307  | 
reduced goal, which needs to be solved by the given tactic. All  | 
|
308  | 
schematic parameters of the goal are imported into the context as  | 
|
309  | 
fixed ones, which may not be instantiated in the sub-proof.  | 
|
| 20491 | 310  | 
|
| 20472 | 311  | 
  \item @{ML Goal.prove}~@{text "ctxt xs As C tac"} states goal @{text
 | 
| 20474 | 312  | 
  "C"} in the context augmented by fixed variables @{text "xs"} and
 | 
313  | 
  assumptions @{text "As"}, and applies tactic @{text "tac"} to solve
 | 
|
314  | 
it. The latter may depend on the local assumptions being presented  | 
|
315  | 
as facts. The result is in HHF normal form.  | 
|
| 18537 | 316  | 
|
| 20472 | 317  | 
  \item @{ML Goal.prove_multi} is simular to @{ML Goal.prove}, but
 | 
| 20491 | 318  | 
states several conclusions simultaneously. The goal is encoded by  | 
| 21827 | 319  | 
  means of Pure conjunction; @{ML Goal.conjunction_tac} will turn this
 | 
320  | 
into a collection of individual subgoals.  | 
|
| 20472 | 321  | 
|
| 20491 | 322  | 
  \item @{ML Obtain.result}~@{text "tac thms ctxt"} eliminates the
 | 
323  | 
given facts using a tactic, which results in additional fixed  | 
|
324  | 
variables and assumptions in the context. Final results need to be  | 
|
325  | 
exported explicitly.  | 
|
| 20472 | 326  | 
|
327  | 
  \end{description}
 | 
|
328  | 
*}  | 
|
| 
30270
 
61811c9224a6
dummy changes to produce a new changeset of these files;
 
wenzelm 
parents: 
29761 
diff
changeset
 | 
329  | 
|
| 18537 | 330  | 
end  |