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