18537
|
1 |
|
|
2 |
(* $Id$ *)
|
|
3 |
|
|
4 |
theory "proof" imports base begin
|
|
5 |
|
20451
|
6 |
chapter {* Structured proofs *}
|
18537
|
7 |
|
20460
|
8 |
section {* Variables *}
|
20026
|
9 |
|
20064
|
10 |
text FIXME
|
|
11 |
|
20026
|
12 |
text %mlref {*
|
|
13 |
\begin{mldecls}
|
|
14 |
@{index_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\
|
20064
|
15 |
@{index_ML Variable.add_fixes: "string list -> Proof.context -> string list * Proof.context"} \\
|
20459
|
16 |
@{index_ML Variable.import: "bool ->
|
|
17 |
thm list -> Proof.context -> ((ctyp list * cterm list) * thm list) * Proof.context"} \\
|
20026
|
18 |
@{index_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\
|
|
19 |
@{index_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\
|
|
20 |
\end{mldecls}
|
|
21 |
|
|
22 |
\begin{description}
|
|
23 |
|
20064
|
24 |
\item @{ML Variable.declare_term}~@{text "t ctxt"} declares term
|
|
25 |
@{text "t"} to belong to the context. This fixes free type
|
|
26 |
variables, but not term variables. Constraints for type and term
|
|
27 |
variables are declared uniformly.
|
|
28 |
|
|
29 |
\item @{ML Variable.add_fixes}~@{text "xs ctxt"} fixes term
|
|
30 |
variables @{text "xs"} and returns the internal names of the
|
|
31 |
resulting Skolem constants. Note that term fixes refer to
|
|
32 |
\emph{all} type instances that may occur in the future.
|
|
33 |
|
|
34 |
\item @{ML Variable.invent_fixes} is similar to @{ML
|
|
35 |
Variable.add_fixes}, but the given names merely act as hints for
|
|
36 |
internal fixes produced here.
|
20026
|
37 |
|
20064
|
38 |
\item @{ML Variable.import}~@{text "open ths ctxt"} augments the
|
|
39 |
context by new fixes for the schematic type and term variables
|
|
40 |
occurring in @{text "ths"}. The @{text "open"} flag indicates
|
|
41 |
whether the fixed names should be accessible to the user, otherwise
|
|
42 |
internal names are chosen.
|
20026
|
43 |
|
20064
|
44 |
\item @{ML Variable.export}~@{text "inner outer ths"} generalizes
|
|
45 |
fixed type and term variables in @{text "ths"} according to the
|
|
46 |
difference of the @{text "inner"} and @{text "outer"} context. Note
|
|
47 |
that type variables occurring in term variables are still fixed.
|
|
48 |
|
|
49 |
@{ML Variable.export} essentially reverses the effect of @{ML
|
|
50 |
Variable.import} (up to renaming of schematic variables.
|
20026
|
51 |
|
20064
|
52 |
\item @{ML Variable.polymorphic}~@{text "ctxt ts"} generalizes type
|
|
53 |
variables in @{text "ts"} as far as possible, even those occurring
|
|
54 |
in fixed term variables. This operation essentially reverses the
|
|
55 |
default policy of type-inference to introduce local polymorphism as
|
|
56 |
fixed types.
|
20026
|
57 |
|
|
58 |
\end{description}
|
|
59 |
*}
|
|
60 |
|
18537
|
61 |
text FIXME
|
|
62 |
|
20451
|
63 |
|
|
64 |
section {* Assumptions *}
|
|
65 |
|
20458
|
66 |
text {*
|
|
67 |
An \emph{assumption} is a proposition that it is postulated in the
|
|
68 |
current context. Local conclusions may use assumptions as
|
|
69 |
additional facts, but this imposes implicit hypotheses that weaken
|
|
70 |
the overall statement.
|
|
71 |
|
|
72 |
Assumptions are restricted to fixed non-schematic statements, all
|
|
73 |
generality needs to be expressed by explicit quantifiers.
|
|
74 |
Nevertheless, the result will be in HHF normal form with outermost
|
|
75 |
quantifiers stripped. For example, by assuming @{text "\<And>x :: \<alpha>. P
|
|
76 |
x"} we get @{text "\<And>x :: \<alpha>. P x \<turnstile> P ?x"} for arbitrary @{text "?x"}
|
|
77 |
of the fixed type @{text "\<alpha>"}. Local derivations accumulate more
|
|
78 |
and more explicit references to hypotheses: @{text "A\<^isub>1, \<dots>,
|
|
79 |
A\<^isub>n \<turnstile> B"} where @{text "A\<^isub>1, \<dots>, A\<^isub>n"} needs to
|
|
80 |
be covered by the assumptions of the current context.
|
|
81 |
|
20459
|
82 |
\medskip The @{text "add_assms"} operation augments the context by
|
|
83 |
local assumptions, which are parameterized by an arbitrary @{text
|
|
84 |
"export"} rule (see below).
|
20458
|
85 |
|
|
86 |
The @{text "export"} operation moves facts from a (larger) inner
|
|
87 |
context into a (smaller) outer context, by discharging the
|
|
88 |
difference of the assumptions as specified by the associated export
|
|
89 |
rules. Note that the discharged portion is determined by the
|
20459
|
90 |
difference contexts, not the facts being exported! There is a
|
|
91 |
separate flag to indicate a goal context, where the result is meant
|
|
92 |
to refine an enclosing sub-goal of a structured proof state (cf.\
|
|
93 |
\secref{sec:isar-proof-state}).
|
20458
|
94 |
|
|
95 |
\medskip The most basic export rule discharges assumptions directly
|
|
96 |
by means of the @{text "\<Longrightarrow>"} introduction rule:
|
|
97 |
\[
|
|
98 |
\infer[(@{text "\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
|
|
99 |
\]
|
|
100 |
|
|
101 |
The variant for goal refinements marks the newly introduced
|
|
102 |
premises, which causes the builtin goal refinement scheme of Isar to
|
|
103 |
enforce unification with local premises within the goal:
|
|
104 |
\[
|
|
105 |
\infer[(@{text "#\<Longrightarrow>_intro"})]{@{text "\<Gamma> \\ A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
|
|
106 |
\]
|
|
107 |
|
20459
|
108 |
\medskip Alternative assumptions may perform arbitrary
|
|
109 |
transformations on export, as long as a particular portion of
|
|
110 |
hypotheses is removed from the given facts. For example, a local
|
|
111 |
definition works by fixing @{text "x"} and assuming @{text "x \<equiv> t"},
|
|
112 |
with the following export rule to reverse the effect:
|
20458
|
113 |
\[
|
|
114 |
\infer{@{text "\<Gamma> \\ x \<equiv> t \<turnstile> B t"}}{@{text "\<Gamma> \<turnstile> B x"}}
|
|
115 |
\]
|
|
116 |
|
|
117 |
\medskip The general concept supports block-structured reasoning
|
|
118 |
nicely, with arbitrary mechanisms for introducing local assumptions.
|
|
119 |
The common reasoning pattern is as follows:
|
|
120 |
|
|
121 |
\medskip
|
|
122 |
\begin{tabular}{l}
|
20459
|
123 |
@{text "add_assms e\<^isub>1 A\<^isub>1"} \\
|
20458
|
124 |
@{text "\<dots>"} \\
|
20459
|
125 |
@{text "add_assms e\<^isub>n A\<^isub>n"} \\
|
20458
|
126 |
@{text "export"} \\
|
|
127 |
\end{tabular}
|
|
128 |
\medskip
|
|
129 |
|
|
130 |
\noindent The final @{text "export"} will turn any fact @{text
|
|
131 |
"A\<^isub>1, \<dots>, A\<^isub>n \<turnstile> B"} into some @{text "\<turnstile> B'"}, by
|
|
132 |
applying the export rules @{text "e\<^isub>1, \<dots>, e\<^isub>n"}
|
|
133 |
inside-out.
|
|
134 |
*}
|
|
135 |
|
|
136 |
text %mlref {*
|
|
137 |
\begin{mldecls}
|
|
138 |
@{index_ML_type Assumption.export} \\
|
|
139 |
@{index_ML Assumption.assume: "cterm -> thm"} \\
|
|
140 |
@{index_ML Assumption.add_assms:
|
20459
|
141 |
"Assumption.export ->
|
|
142 |
cterm list -> Proof.context -> thm list * Proof.context"} \\
|
|
143 |
@{index_ML Assumption.add_assumes: "
|
|
144 |
cterm list -> Proof.context -> thm list * Proof.context"} \\
|
20458
|
145 |
@{index_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\
|
|
146 |
\end{mldecls}
|
|
147 |
|
|
148 |
\begin{description}
|
|
149 |
|
20459
|
150 |
\item @{ML_type Assumption.export} represents arbitrary export
|
|
151 |
rules, which is any function of type @{ML_type "bool -> cterm list -> thm -> thm"},
|
|
152 |
where the @{ML_type "bool"} indicates goal mode, and the @{ML_type
|
|
153 |
"cterm list"} the collection of assumptions to be discharged
|
|
154 |
simultaneously.
|
20458
|
155 |
|
20459
|
156 |
\item @{ML Assumption.assume}~@{text "A"} turns proposition @{text
|
|
157 |
"A"} into a raw assumption @{text "A \<turnstile> A'"}, where the conclusion
|
|
158 |
@{text "A'"} is in HHF normal form.
|
20458
|
159 |
|
|
160 |
\item @{ML Assumption.add_assms}~@{text "e As"} augments the context
|
20459
|
161 |
by assumptions @{text "As"} with export rule @{text "e"}. The
|
|
162 |
resulting facts are hypothetical theorems as produced by @{ML
|
|
163 |
Assumption.assume}.
|
|
164 |
|
|
165 |
\item @{ML Assumption.add_assumes}~@{text "As"} is a special case of
|
|
166 |
@{ML Assumption.add_assms} where the export rule performs @{text
|
|
167 |
"\<Longrightarrow>_intro"} or @{text "#\<Longrightarrow>_intro"}, depending on goal mode.
|
20458
|
168 |
|
|
169 |
\item @{ML Assumption.export}~@{text "is_goal inner outer th"}
|
|
170 |
exports result @{text "th"} from the the @{text "inner"} context
|
20459
|
171 |
back into the @{text "outer"} one; @{text "is_goal = true"} means
|
|
172 |
this is a goal context. The result is in HHF normal form. Note
|
|
173 |
that @{ML "ProofContext.export"} combines @{ML "Variable.export"}
|
|
174 |
and @{ML "Assumption.export"} in the canonical way.
|
20458
|
175 |
|
|
176 |
\end{description}
|
|
177 |
*}
|
20451
|
178 |
|
|
179 |
|
|
180 |
section {* Conclusions *}
|
|
181 |
|
|
182 |
text FIXME
|
|
183 |
|
|
184 |
|
20452
|
185 |
section {* Proof states \label{sec:isar-proof-state} *}
|
18537
|
186 |
|
|
187 |
text {*
|
|
188 |
FIXME
|
|
189 |
|
|
190 |
\glossary{Proof state}{The whole configuration of a structured proof,
|
|
191 |
consisting of a \seeglossary{proof context} and an optional
|
|
192 |
\seeglossary{structured goal}. Internally, an Isar proof state is
|
|
193 |
organized as a stack to accomodate block structure of proof texts.
|
|
194 |
For historical reasons, a low-level \seeglossary{tactical goal} is
|
|
195 |
occasionally called ``proof state'' as well.}
|
|
196 |
|
|
197 |
\glossary{Structured goal}{FIXME}
|
|
198 |
|
|
199 |
\glossary{Goal}{See \seeglossary{tactical goal} or \seeglossary{structured goal}. \norefpage}
|
|
200 |
|
|
201 |
|
|
202 |
*}
|
|
203 |
|
20451
|
204 |
section {* Proof methods *}
|
18537
|
205 |
|
|
206 |
text FIXME
|
|
207 |
|
|
208 |
section {* Attributes *}
|
|
209 |
|
20451
|
210 |
text "FIXME ?!"
|
18537
|
211 |
|
|
212 |
end
|