author | wenzelm |
Thu, 22 Oct 2015 23:01:49 +0200 | |
changeset 61506 | 436b7fe89cdc |
parent 61493 | 0debd22f0c0e |
child 61580 | c49a8ebd30cc |
permissions | -rw-r--r-- |
29755 | 1 |
theory Proof |
2 |
imports Base |
|
3 |
begin |
|
18537 | 4 |
|
58618 | 5 |
chapter \<open>Structured proofs\<close> |
18537 | 6 |
|
58618 | 7 |
section \<open>Variables \label{sec:variables}\<close> |
20026 | 8 |
|
58618 | 9 |
text \<open> |
61493 | 10 |
Any variable that is not explicitly bound by \<open>\<lambda>\<close>-abstraction |
20470 | 11 |
is considered as ``free''. Logically, free variables act like |
61493 | 12 |
outermost universal quantification at the sequent level: \<open>A\<^sub>1(x), \<dots>, A\<^sub>n(x) \<turnstile> B(x)\<close> means that the result |
13 |
holds \<^emph>\<open>for all\<close> values of \<open>x\<close>. Free variables for |
|
14 |
terms (not types) can be fully internalized into the logic: \<open>\<turnstile> B(x)\<close> and \<open>\<turnstile> \<And>x. B(x)\<close> are interchangeable, provided |
|
15 |
that \<open>x\<close> does not occur elsewhere in the context. |
|
16 |
Inspecting \<open>\<turnstile> \<And>x. B(x)\<close> more closely, we see that inside the |
|
17 |
quantifier, \<open>x\<close> is essentially ``arbitrary, but fixed'', |
|
20470 | 18 |
while from outside it appears as a place-holder for instantiation |
61493 | 19 |
(thanks to \<open>\<And>\<close> elimination). |
20470 | 20 |
|
20474 | 21 |
The Pure logic represents the idea of variables being either inside |
22 |
or outside the current scope by providing separate syntactic |
|
61493 | 23 |
categories for \<^emph>\<open>fixed variables\<close> (e.g.\ \<open>x\<close>) vs.\ |
24 |
\<^emph>\<open>schematic variables\<close> (e.g.\ \<open>?x\<close>). Incidently, a |
|
25 |
universal result \<open>\<turnstile> \<And>x. B(x)\<close> has the HHF normal form \<open>\<turnstile> B(?x)\<close>, which represents its generality without requiring an |
|
34930 | 26 |
explicit quantifier. The same principle works for type variables: |
61493 | 27 |
\<open>\<turnstile> B(?\<alpha>)\<close> represents the idea of ``\<open>\<turnstile> \<forall>\<alpha>. B(\<alpha>)\<close>'' |
34930 | 28 |
without demanding a truly polymorphic framework. |
20470 | 29 |
|
61416 | 30 |
\<^medskip> |
31 |
Additional care is required to treat type variables in a |
|
20470 | 32 |
way that facilitates type-inference. In principle, term variables |
33 |
depend on type variables, which means that type variables would have |
|
34 |
to be declared first. For example, a raw type-theoretic framework |
|
35 |
would demand the context to be constructed in stages as follows: |
|
61493 | 36 |
\<open>\<Gamma> = \<alpha>: type, x: \<alpha>, a: A(x\<^sub>\<alpha>)\<close>. |
20470 | 37 |
|
38 |
We allow a slightly less formalistic mode of operation: term |
|
61493 | 39 |
variables \<open>x\<close> are fixed without specifying a type yet |
61477 | 40 |
(essentially \<^emph>\<open>all\<close> potential occurrences of some instance |
61493 | 41 |
\<open>x\<^sub>\<tau>\<close> are fixed); the first occurrence of \<open>x\<close> |
20474 | 42 |
within a specific term assigns its most general type, which is then |
43 |
maintained consistently in the context. The above example becomes |
|
61493 | 44 |
\<open>\<Gamma> = x: term, \<alpha>: type, A(x\<^sub>\<alpha>)\<close>, where type \<open>\<alpha>\<close> is fixed \<^emph>\<open>after\<close> term \<open>x\<close>, and the constraint |
45 |
\<open>x :: \<alpha>\<close> is an implicit consequence of the occurrence of |
|
46 |
\<open>x\<^sub>\<alpha>\<close> in the subsequent proposition. |
|
20470 | 47 |
|
48 |
This twist of dependencies is also accommodated by the reverse |
|
49 |
operation of exporting results from a context: a type variable |
|
61493 | 50 |
\<open>\<alpha>\<close> is considered fixed as long as it occurs in some fixed |
51 |
term variable of the context. For example, exporting \<open>x: |
|
52 |
term, \<alpha>: type \<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> produces in the first step \<open>x: term |
|
53 |
\<turnstile> x\<^sub>\<alpha> \<equiv> x\<^sub>\<alpha>\<close> for fixed \<open>\<alpha>\<close>, and only in the second step |
|
54 |
\<open>\<turnstile> ?x\<^sub>?\<^sub>\<alpha> \<equiv> ?x\<^sub>?\<^sub>\<alpha>\<close> for schematic \<open>?x\<close> and \<open>?\<alpha>\<close>. |
|
39841 | 55 |
The following Isar source text illustrates this scenario. |
58618 | 56 |
\<close> |
20470 | 57 |
|
40964 | 58 |
notepad |
59 |
begin |
|
39841 | 60 |
{ |
61493 | 61 |
fix x -- \<open>all potential occurrences of some \<open>x::\<tau>\<close> are fixed\<close> |
39841 | 62 |
{ |
58618 | 63 |
have "x::'a \<equiv> x" -- \<open>implicit type assignment by concrete occurrence\<close> |
39841 | 64 |
by (rule reflexive) |
65 |
} |
|
61493 | 66 |
thm this -- \<open>result still with fixed type \<open>'a\<close>\<close> |
39841 | 67 |
} |
61493 | 68 |
thm this -- \<open>fully general result for arbitrary \<open>?x::?'a\<close>\<close> |
40964 | 69 |
end |
39841 | 70 |
|
58618 | 71 |
text \<open>The Isabelle/Isar proof context manages the details of term |
39861
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents:
39853
diff
changeset
|
72 |
vs.\ type variables, with high-level principles for moving the |
20474 | 73 |
frontier between fixed and schematic variables. |
74 |
||
61493 | 75 |
The \<open>add_fixes\<close> operation explicitly declares fixed |
76 |
variables; the \<open>declare_term\<close> operation absorbs a term into |
|
20474 | 77 |
a context by fixing new type variables and adding syntactic |
78 |
constraints. |
|
20470 | 79 |
|
61493 | 80 |
The \<open>export\<close> operation is able to perform the main work of |
20474 | 81 |
generalizing term and type variables as sketched above, assuming |
82 |
that fixing variables and terms have been declared properly. |
|
83 |
||
61493 | 84 |
There \<open>import\<close> operation makes a generalized fact a genuine |
20474 | 85 |
part of the context, by inventing fixed variables for the schematic |
61493 | 86 |
ones. The effect can be reversed by using \<open>export\<close> later, |
20474 | 87 |
potentially with an extended context; the result is equivalent to |
88 |
the original modulo renaming of schematic variables. |
|
20470 | 89 |
|
61493 | 90 |
The \<open>focus\<close> operation provides a variant of \<open>import\<close> |
91 |
for nested propositions (with explicit quantification): \<open>\<And>x\<^sub>1 \<dots> x\<^sub>n. B(x\<^sub>1, \<dots>, x\<^sub>n)\<close> is |
|
92 |
decomposed by inventing fixed variables \<open>x\<^sub>1, \<dots>, |
|
93 |
x\<^sub>n\<close> for the body. |
|
58618 | 94 |
\<close> |
20064 | 95 |
|
58618 | 96 |
text %mlref \<open> |
20026 | 97 |
\begin{mldecls} |
20474 | 98 |
@{index_ML Variable.add_fixes: " |
99 |
string list -> Proof.context -> string list * Proof.context"} \\ |
|
20797
c1f0bc7e7d80
renamed Variable.invent_fixes to Variable.variant_fixes;
wenzelm
parents:
20547
diff
changeset
|
100 |
@{index_ML Variable.variant_fixes: " |
20474 | 101 |
string list -> Proof.context -> string list * Proof.context"} \\ |
20026 | 102 |
@{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\ |
20470 | 103 |
@{index_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\ |
104 |
@{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\ |
|
105 |
@{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\ |
|
31794
71af1fd6a5e4
renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents:
30272
diff
changeset
|
106 |
@{index_ML Variable.import: "bool -> thm list -> Proof.context -> |
60642
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
wenzelm
parents:
59902
diff
changeset
|
107 |
((((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list) * thm list) * Proof.context"} \\ |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
108 |
@{index_ML Variable.focus: "binding list option -> term -> Proof.context -> |
42509 | 109 |
((string * (string * typ)) list * term) * Proof.context"} \\ |
20026 | 110 |
\end{mldecls} |
111 |
||
61493 | 112 |
\<^descr> @{ML Variable.add_fixes}~\<open>xs ctxt\<close> fixes term |
113 |
variables \<open>xs\<close>, returning the resulting internal names. By |
|
20470 | 114 |
default, the internal representation coincides with the external |
20474 | 115 |
one, which also means that the given variables must not be fixed |
116 |
already. There is a different policy within a local proof body: the |
|
117 |
given names are just hints for newly invented Skolem variables. |
|
20064 | 118 |
|
61439 | 119 |
\<^descr> @{ML Variable.variant_fixes} is similar to @{ML |
20470 | 120 |
Variable.add_fixes}, but always produces fresh variants of the given |
20474 | 121 |
names. |
20470 | 122 |
|
61493 | 123 |
\<^descr> @{ML Variable.declare_term}~\<open>t ctxt\<close> declares term |
124 |
\<open>t\<close> to belong to the context. This automatically fixes new |
|
20470 | 125 |
type variables, but not term variables. Syntactic constraints for |
20474 | 126 |
type and term variables are declared uniformly, though. |
20470 | 127 |
|
61493 | 128 |
\<^descr> @{ML Variable.declare_constraints}~\<open>t ctxt\<close> declares |
129 |
syntactic constraints from term \<open>t\<close>, without making it part |
|
20474 | 130 |
of the context yet. |
20026 | 131 |
|
61493 | 132 |
\<^descr> @{ML Variable.export}~\<open>inner outer thms\<close> generalizes |
133 |
fixed type and term variables in \<open>thms\<close> according to the |
|
134 |
difference of the \<open>inner\<close> and \<open>outer\<close> context, |
|
20470 | 135 |
following the principles sketched above. |
136 |
||
61493 | 137 |
\<^descr> @{ML Variable.polymorphic}~\<open>ctxt ts\<close> generalizes type |
138 |
variables in \<open>ts\<close> as far as possible, even those occurring |
|
20470 | 139 |
in fixed term variables. The default policy of type-inference is to |
20474 | 140 |
fix newly introduced type variables, which is essentially reversed |
141 |
with @{ML Variable.polymorphic}: here the given terms are detached |
|
142 |
from the context as far as possible. |
|
20470 | 143 |
|
61493 | 144 |
\<^descr> @{ML Variable.import}~\<open>open thms ctxt\<close> invents fixed |
145 |
type and term variables for the schematic ones occurring in \<open>thms\<close>. The \<open>open\<close> flag indicates whether the fixed names |
|
20474 | 146 |
should be accessible to the user, otherwise newly introduced names |
147 |
are marked as ``internal'' (\secref{sec:names}). |
|
20026 | 148 |
|
61493 | 149 |
\<^descr> @{ML Variable.focus}~\<open>bindings B\<close> decomposes the outermost \<open>\<And>\<close> prefix of proposition \<open>B\<close>, using the given name bindings. |
58618 | 150 |
\<close> |
20026 | 151 |
|
58618 | 152 |
text %mlex \<open>The following example shows how to work with fixed term |
153 |
and type parameters and with type-inference.\<close> |
|
34932 | 154 |
|
59902 | 155 |
ML_val \<open> |
34932 | 156 |
(*static compile-time context -- for testing only*) |
157 |
val ctxt0 = @{context}; |
|
158 |
||
159 |
(*locally fixed parameters -- no type assignment yet*) |
|
160 |
val ([x, y], ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y"]; |
|
161 |
||
162 |
(*t1: most general fixed type; t1': most general arbitrary type*) |
|
163 |
val t1 = Syntax.read_term ctxt1 "x"; |
|
164 |
val t1' = singleton (Variable.polymorphic ctxt1) t1; |
|
165 |
||
166 |
(*term u enforces specific type assignment*) |
|
39846 | 167 |
val u = Syntax.read_term ctxt1 "(x::nat) \<equiv> y"; |
34932 | 168 |
|
169 |
(*official declaration of u -- propagates constraints etc.*) |
|
170 |
val ctxt2 = ctxt1 |> Variable.declare_term u; |
|
39846 | 171 |
val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*) |
58618 | 172 |
\<close> |
34932 | 173 |
|
58618 | 174 |
text \<open>In the above example, the starting context is derived from the |
40126 | 175 |
toplevel theory, which means that fixed variables are internalized |
61493 | 176 |
literally: \<open>x\<close> is mapped again to \<open>x\<close>, and |
40126 | 177 |
attempting to fix it again in the subsequent context is an error. |
178 |
Alternatively, fixed parameters can be renamed explicitly as |
|
58618 | 179 |
follows:\<close> |
34932 | 180 |
|
59902 | 181 |
ML_val \<open> |
34932 | 182 |
val ctxt0 = @{context}; |
183 |
val ([x1, x2, x3], ctxt1) = |
|
184 |
ctxt0 |> Variable.variant_fixes ["x", "x", "x"]; |
|
58618 | 185 |
\<close> |
34932 | 186 |
|
58618 | 187 |
text \<open>The following ML code can now work with the invented names of |
61493 | 188 |
\<open>x1\<close>, \<open>x2\<close>, \<open>x3\<close>, without depending on |
39861
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents:
39853
diff
changeset
|
189 |
the details on the system policy for introducing these variants. |
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents:
39853
diff
changeset
|
190 |
Recall that within a proof body the system always invents fresh |
58618 | 191 |
``Skolem constants'', e.g.\ as follows:\<close> |
34932 | 192 |
|
40964 | 193 |
notepad |
194 |
begin |
|
58728 | 195 |
ML_prf %"ML" |
196 |
\<open>val ctxt0 = @{context}; |
|
34932 | 197 |
|
198 |
val ([x1], ctxt1) = ctxt0 |> Variable.add_fixes ["x"]; |
|
199 |
val ([x2], ctxt2) = ctxt1 |> Variable.add_fixes ["x"]; |
|
200 |
val ([x3], ctxt3) = ctxt2 |> Variable.add_fixes ["x"]; |
|
201 |
||
202 |
val ([y1, y2], ctxt4) = |
|
58728 | 203 |
ctxt3 |> Variable.variant_fixes ["y", "y"];\<close> |
40964 | 204 |
end |
34932 | 205 |
|
58618 | 206 |
text \<open>In this situation @{ML Variable.add_fixes} and @{ML |
34932 | 207 |
Variable.variant_fixes} are very similar, but identical name |
208 |
proposals given in a row are only accepted by the second version. |
|
58618 | 209 |
\<close> |
34932 | 210 |
|
18537 | 211 |
|
58618 | 212 |
section \<open>Assumptions \label{sec:assumptions}\<close> |
20451 | 213 |
|
58618 | 214 |
text \<open> |
61477 | 215 |
An \<^emph>\<open>assumption\<close> is a proposition that it is postulated in the |
20458 | 216 |
current context. Local conclusions may use assumptions as |
217 |
additional facts, but this imposes implicit hypotheses that weaken |
|
218 |
the overall statement. |
|
219 |
||
20474 | 220 |
Assumptions are restricted to fixed non-schematic statements, i.e.\ |
221 |
all generality needs to be expressed by explicit quantifiers. |
|
20458 | 222 |
Nevertheless, the result will be in HHF normal form with outermost |
61493 | 223 |
quantifiers stripped. For example, by assuming \<open>\<And>x :: \<alpha>. P |
224 |
x\<close> we get \<open>\<And>x :: \<alpha>. P x \<turnstile> P ?x\<close> for schematic \<open>?x\<close> |
|
225 |
of fixed type \<open>\<alpha>\<close>. Local derivations accumulate more and |
|
226 |
more explicit references to hypotheses: \<open>A\<^sub>1, \<dots>, |
|
227 |
A\<^sub>n \<turnstile> B\<close> where \<open>A\<^sub>1, \<dots>, A\<^sub>n\<close> needs to |
|
20458 | 228 |
be covered by the assumptions of the current context. |
229 |
||
61416 | 230 |
\<^medskip> |
61493 | 231 |
The \<open>add_assms\<close> operation augments the context by |
232 |
local assumptions, which are parameterized by an arbitrary \<open>export\<close> rule (see below). |
|
20458 | 233 |
|
61493 | 234 |
The \<open>export\<close> operation moves facts from a (larger) inner |
20458 | 235 |
context into a (smaller) outer context, by discharging the |
236 |
difference of the assumptions as specified by the associated export |
|
237 |
rules. Note that the discharged portion is determined by the |
|
34930 | 238 |
difference of contexts, not the facts being exported! There is a |
20459 | 239 |
separate flag to indicate a goal context, where the result is meant |
29760 | 240 |
to refine an enclosing sub-goal of a structured proof state. |
20458 | 241 |
|
61416 | 242 |
\<^medskip> |
243 |
The most basic export rule discharges assumptions directly |
|
61493 | 244 |
by means of the \<open>\<Longrightarrow>\<close> introduction rule: |
20458 | 245 |
\[ |
61493 | 246 |
\infer[(\<open>\<Longrightarrow>\<hyphen>intro\<close>)]{\<open>\<Gamma> - A \<turnstile> A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>} |
20458 | 247 |
\] |
248 |
||
249 |
The variant for goal refinements marks the newly introduced |
|
20474 | 250 |
premises, which causes the canonical Isar goal refinement scheme to |
20458 | 251 |
enforce unification with local premises within the goal: |
252 |
\[ |
|
61493 | 253 |
\infer[(\<open>#\<Longrightarrow>\<hyphen>intro\<close>)]{\<open>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B\<close>}{\<open>\<Gamma> \<turnstile> B\<close>} |
20458 | 254 |
\] |
255 |
||
61416 | 256 |
\<^medskip> |
257 |
Alternative versions of assumptions may perform arbitrary |
|
20474 | 258 |
transformations on export, as long as the corresponding portion of |
20459 | 259 |
hypotheses is removed from the given facts. For example, a local |
61493 | 260 |
definition works by fixing \<open>x\<close> and assuming \<open>x \<equiv> t\<close>, |
20459 | 261 |
with the following export rule to reverse the effect: |
20458 | 262 |
\[ |
61493 | 263 |
\infer[(\<open>\<equiv>\<hyphen>expand\<close>)]{\<open>\<Gamma> - (x \<equiv> t) \<turnstile> B t\<close>}{\<open>\<Gamma> \<turnstile> B x\<close>} |
20458 | 264 |
\] |
61493 | 265 |
This works, because the assumption \<open>x \<equiv> t\<close> was introduced in |
266 |
a context with \<open>x\<close> being fresh, so \<open>x\<close> does not |
|
267 |
occur in \<open>\<Gamma>\<close> here. |
|
58618 | 268 |
\<close> |
20458 | 269 |
|
58618 | 270 |
text %mlref \<open> |
20458 | 271 |
\begin{mldecls} |
272 |
@{index_ML_type Assumption.export} \\ |
|
54883
dd04a8b654fc
proper context for norm_hhf and derived operations;
wenzelm
parents:
53015
diff
changeset
|
273 |
@{index_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\ |
20458 | 274 |
@{index_ML Assumption.add_assms: |
20459 | 275 |
"Assumption.export -> |
276 |
cterm list -> Proof.context -> thm list * Proof.context"} \\ |
|
277 |
@{index_ML Assumption.add_assumes: " |
|
278 |
cterm list -> Proof.context -> thm list * Proof.context"} \\ |
|
20458 | 279 |
@{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\ |
280 |
\end{mldecls} |
|
281 |
||
61439 | 282 |
\<^descr> Type @{ML_type Assumption.export} represents arbitrary export |
39864 | 283 |
rules, which is any function of type @{ML_type "bool -> cterm list |
284 |
-> thm -> thm"}, where the @{ML_type "bool"} indicates goal mode, |
|
285 |
and the @{ML_type "cterm list"} the collection of assumptions to be |
|
286 |
discharged simultaneously. |
|
20458 | 287 |
|
61493 | 288 |
\<^descr> @{ML Assumption.assume}~\<open>ctxt A\<close> turns proposition \<open>A\<close> into a primitive assumption \<open>A \<turnstile> A'\<close>, where the |
289 |
conclusion \<open>A'\<close> is in HHF normal form. |
|
20458 | 290 |
|
61493 | 291 |
\<^descr> @{ML Assumption.add_assms}~\<open>r As\<close> augments the context |
292 |
by assumptions \<open>As\<close> with export rule \<open>r\<close>. The |
|
20474 | 293 |
resulting facts are hypothetical theorems as produced by the raw |
294 |
@{ML Assumption.assume}. |
|
20459 | 295 |
|
61493 | 296 |
\<^descr> @{ML Assumption.add_assumes}~\<open>As\<close> is a special case of |
297 |
@{ML Assumption.add_assms} where the export rule performs \<open>\<Longrightarrow>\<hyphen>intro\<close> or \<open>#\<Longrightarrow>\<hyphen>intro\<close>, depending on goal |
|
34930 | 298 |
mode. |
20458 | 299 |
|
61493 | 300 |
\<^descr> @{ML Assumption.export}~\<open>is_goal inner outer thm\<close> |
301 |
exports result \<open>thm\<close> from the the \<open>inner\<close> context |
|
302 |
back into the \<open>outer\<close> one; \<open>is_goal = true\<close> means |
|
20459 | 303 |
this is a goal context. The result is in HHF normal form. Note |
42361 | 304 |
that @{ML "Proof_Context.export"} combines @{ML "Variable.export"} |
20459 | 305 |
and @{ML "Assumption.export"} in the canonical way. |
58618 | 306 |
\<close> |
20451 | 307 |
|
58618 | 308 |
text %mlex \<open>The following example demonstrates how rules can be |
34932 | 309 |
derived by building up a context of assumptions first, and exporting |
310 |
some local fact afterwards. We refer to @{theory Pure} equality |
|
311 |
here for testing purposes. |
|
58618 | 312 |
\<close> |
34932 | 313 |
|
59902 | 314 |
ML_val \<open> |
34932 | 315 |
(*static compile-time context -- for testing only*) |
316 |
val ctxt0 = @{context}; |
|
317 |
||
318 |
val ([eq], ctxt1) = |
|
319 |
ctxt0 |> Assumption.add_assumes [@{cprop "x \<equiv> y"}]; |
|
320 |
val eq' = Thm.symmetric eq; |
|
321 |
||
322 |
(*back to original context -- discharges assumption*) |
|
323 |
val r = Assumption.export false ctxt1 ctxt0 eq'; |
|
58618 | 324 |
\<close> |
34932 | 325 |
|
58618 | 326 |
text \<open>Note that the variables of the resulting rule are not |
39861
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents:
39853
diff
changeset
|
327 |
generalized. This would have required to fix them properly in the |
b8d89db3e238
use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents:
39853
diff
changeset
|
328 |
context beforehand, and export wrt.\ variables afterwards (cf.\ @{ML |
58618 | 329 |
Variable.export} or the combined @{ML "Proof_Context.export"}).\<close> |
34932 | 330 |
|
20451 | 331 |
|
58618 | 332 |
section \<open>Structured goals and results \label{sec:struct-goals}\<close> |
20451 | 333 |
|
58618 | 334 |
text \<open> |
20472 | 335 |
Local results are established by monotonic reasoning from facts |
20474 | 336 |
within a context. This allows common combinations of theorems, |
61493 | 337 |
e.g.\ via \<open>\<And>/\<Longrightarrow>\<close> elimination, resolution rules, or equational |
20474 | 338 |
reasoning, see \secref{sec:thms}. Unaccounted context manipulations |
61493 | 339 |
should be avoided, notably raw \<open>\<And>/\<Longrightarrow>\<close> introduction or ad-hoc |
20472 | 340 |
references to free variables or assumptions not present in the proof |
341 |
context. |
|
18537 | 342 |
|
61416 | 343 |
\<^medskip> |
61493 | 344 |
The \<open>SUBPROOF\<close> combinator allows to structure a |
20491 | 345 |
tactical proof recursively by decomposing a selected sub-goal: |
61493 | 346 |
\<open>(\<And>x. A(x) \<Longrightarrow> B(x)) \<Longrightarrow> \<dots>\<close> is turned into \<open>B(x) \<Longrightarrow> \<dots>\<close> |
347 |
after fixing \<open>x\<close> and assuming \<open>A(x)\<close>. This means |
|
20491 | 348 |
the tactic needs to solve the conclusion, but may use the premise as |
349 |
a local fact, for locally fixed variables. |
|
18537 | 350 |
|
61493 | 351 |
The family of \<open>FOCUS\<close> combinators is similar to \<open>SUBPROOF\<close>, but allows to retain schematic variables and pending |
34930 | 352 |
subgoals in the resulting goal state. |
353 |
||
61493 | 354 |
The \<open>prove\<close> operation provides an interface for structured |
20491 | 355 |
backwards reasoning under program control, with some explicit sanity |
356 |
checks of the result. The goal context can be augmented by |
|
357 |
additional fixed variables (cf.\ \secref{sec:variables}) and |
|
358 |
assumptions (cf.\ \secref{sec:assumptions}), which will be available |
|
359 |
as local facts during the proof and discharged into implications in |
|
360 |
the result. Type and term variables are generalized as usual, |
|
361 |
according to the context. |
|
18537 | 362 |
|
61493 | 363 |
The \<open>obtain\<close> operation produces results by eliminating |
20491 | 364 |
existing facts by means of a given tactic. This acts like a dual |
365 |
conclusion: the proof demonstrates that the context may be augmented |
|
34930 | 366 |
by parameters and assumptions, without affecting any conclusions |
367 |
that do not mention these parameters. See also |
|
58555 | 368 |
@{cite "isabelle-isar-ref"} for the user-level @{command obtain} and |
39851 | 369 |
@{command guess} elements. Final results, which may not refer to |
20491 | 370 |
the parameters in the conclusion, need to exported explicitly into |
58618 | 371 |
the original context.\<close> |
18537 | 372 |
|
58618 | 373 |
text %mlref \<open> |
20472 | 374 |
\begin{mldecls} |
39821 | 375 |
@{index_ML SUBPROOF: "(Subgoal.focus -> tactic) -> |
376 |
Proof.context -> int -> tactic"} \\ |
|
377 |
@{index_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> |
|
378 |
Proof.context -> int -> tactic"} \\ |
|
379 |
@{index_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> |
|
380 |
Proof.context -> int -> tactic"} \\ |
|
381 |
@{index_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> |
|
382 |
Proof.context -> int -> tactic"} \\ |
|
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
383 |
@{index_ML Subgoal.focus: "Proof.context -> int -> binding list option -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
384 |
thm -> Subgoal.focus * thm"} \\ |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
385 |
@{index_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
386 |
thm -> Subgoal.focus * thm"} \\ |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
387 |
@{index_ML Subgoal.focus_params: "Proof.context -> int -> binding list option -> |
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
388 |
thm -> Subgoal.focus * thm"} \\ |
20547 | 389 |
\end{mldecls} |
34930 | 390 |
|
20547 | 391 |
\begin{mldecls} |
20472 | 392 |
@{index_ML Goal.prove: "Proof.context -> string list -> term list -> term -> |
393 |
({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\ |
|
59564
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
394 |
@{index_ML Goal.prove_common: "Proof.context -> int option -> |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
395 |
string list -> term list -> term list -> |
20472 | 396 |
({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\ |
20547 | 397 |
\end{mldecls} |
398 |
\begin{mldecls} |
|
39821 | 399 |
@{index_ML Obtain.result: "(Proof.context -> tactic) -> thm list -> |
400 |
Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ |
|
20472 | 401 |
\end{mldecls} |
18537 | 402 |
|
61493 | 403 |
\<^descr> @{ML SUBPROOF}~\<open>tac ctxt i\<close> decomposes the structure |
29761 | 404 |
of the specified sub-goal, producing an extended context and a |
405 |
reduced goal, which needs to be solved by the given tactic. All |
|
406 |
schematic parameters of the goal are imported into the context as |
|
407 |
fixed ones, which may not be instantiated in the sub-proof. |
|
20491 | 408 |
|
61439 | 409 |
\<^descr> @{ML Subgoal.FOCUS}, @{ML Subgoal.FOCUS_PREMS}, and @{ML |
34930 | 410 |
Subgoal.FOCUS_PARAMS} are similar to @{ML SUBPROOF}, but are |
411 |
slightly more flexible: only the specified parts of the subgoal are |
|
412 |
imported into the context, and the body tactic may introduce new |
|
413 |
subgoals and schematic variables. |
|
414 |
||
61439 | 415 |
\<^descr> @{ML Subgoal.focus}, @{ML Subgoal.focus_prems}, @{ML |
39853 | 416 |
Subgoal.focus_params} extract the focus information from a goal |
417 |
state in the same way as the corresponding tacticals above. This is |
|
418 |
occasionally useful to experiment without writing actual tactics |
|
419 |
yet. |
|
420 |
||
61493 | 421 |
\<^descr> @{ML Goal.prove}~\<open>ctxt xs As C tac\<close> states goal \<open>C\<close> in the context augmented by fixed variables \<open>xs\<close> and |
422 |
assumptions \<open>As\<close>, and applies tactic \<open>tac\<close> to solve |
|
20474 | 423 |
it. The latter may depend on the local assumptions being presented |
424 |
as facts. The result is in HHF normal form. |
|
18537 | 425 |
|
61493 | 426 |
\<^descr> @{ML Goal.prove_common}~\<open>ctxt fork_pri\<close> is the common form |
59564
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
427 |
to state and prove a simultaneous goal statement, where @{ML Goal.prove} |
59567 | 428 |
is a convenient shorthand that is most frequently used in applications. |
59564
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
429 |
|
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
430 |
The given list of simultaneous conclusions is encoded in the goal state by |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
431 |
means of Pure conjunction: @{ML Goal.conjunction_tac} will turn this into |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
432 |
a collection of individual subgoals, but note that the original multi-goal |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
433 |
state is usually required for advanced induction. |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
434 |
|
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
435 |
It is possible to provide an optional priority for a forked proof, |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
436 |
typically @{ML "SOME ~1"}, while @{ML NONE} means the proof is immediate |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
437 |
(sequential) as for @{ML Goal.prove}. Note that a forked proof does not |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
438 |
exhibit any failures in the usual way via exceptions in ML, but |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
439 |
accumulates error situations under the execution id of the running |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
440 |
transaction. Thus the system is able to expose error messages ultimately |
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
wenzelm
parents:
58801
diff
changeset
|
441 |
to the end-user, even though the subsequent ML code misses them. |
20472 | 442 |
|
61493 | 443 |
\<^descr> @{ML Obtain.result}~\<open>tac thms ctxt\<close> eliminates the |
20491 | 444 |
given facts using a tactic, which results in additional fixed |
445 |
variables and assumptions in the context. Final results need to be |
|
446 |
exported explicitly. |
|
58618 | 447 |
\<close> |
30272 | 448 |
|
58618 | 449 |
text %mlex \<open>The following minimal example illustrates how to access |
450 |
the focus information of a structured goal state.\<close> |
|
39853 | 451 |
|
40964 | 452 |
notepad |
453 |
begin |
|
39853 | 454 |
fix A B C :: "'a \<Rightarrow> bool" |
455 |
||
456 |
have "\<And>x. A x \<Longrightarrow> B x \<Longrightarrow> C x" |
|
457 |
ML_val |
|
58728 | 458 |
\<open>val {goal, context = goal_ctxt, ...} = @{Isar.goal}; |
39853 | 459 |
val (focus as {params, asms, concl, ...}, goal') = |
60695
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
wenzelm
parents:
60642
diff
changeset
|
460 |
Subgoal.focus goal_ctxt 1 (SOME [@{binding x}]) goal; |
39853 | 461 |
val [A, B] = #prems focus; |
58728 | 462 |
val [(_, x)] = #params focus;\<close> |
58801 | 463 |
sorry |
464 |
end |
|
39853 | 465 |
|
61416 | 466 |
text \<open> |
467 |
\<^medskip> |
|
468 |
The next example demonstrates forward-elimination in |
|
58618 | 469 |
a local context, using @{ML Obtain.result}.\<close> |
39851 | 470 |
|
40964 | 471 |
notepad |
472 |
begin |
|
39851 | 473 |
assume ex: "\<exists>x. B x" |
474 |
||
58728 | 475 |
ML_prf %"ML" |
476 |
\<open>val ctxt0 = @{context}; |
|
39851 | 477 |
val (([(_, x)], [B]), ctxt1) = ctxt0 |
60754 | 478 |
|> Obtain.result (fn _ => eresolve_tac ctxt0 @{thms exE} 1) [@{thm ex}];\<close> |
58728 | 479 |
ML_prf %"ML" |
480 |
\<open>singleton (Proof_Context.export ctxt1 ctxt0) @{thm refl};\<close> |
|
481 |
ML_prf %"ML" |
|
482 |
\<open>Proof_Context.export ctxt1 ctxt0 [Thm.reflexive x] |
|
483 |
handle ERROR msg => (warning msg; []);\<close> |
|
40964 | 484 |
end |
39851 | 485 |
|
18537 | 486 |
end |