| author | Christian Urban <urbanc@in.tum.de> |
| Thu, 17 Nov 2011 16:33:49 +0000 | |
| changeset 502 | 615780a701b6 |
| parent 501 | f56fc3305a08 |
| child 506 | caa733190454 |
| permissions | -rw-r--r-- |
|
395
2c392f61f400
spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents:
394
diff
changeset
|
1 |
theory Advanced |
| 441 | 2 |
imports Base First_Steps |
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
begin |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
|
|
346
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
5 |
(*<*) |
|
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
6 |
setup{*
|
|
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
7 |
open_file_with_prelude |
|
395
2c392f61f400
spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents:
394
diff
changeset
|
8 |
"Advanced_Code.thy" |
| 441 | 9 |
["theory Advanced", "imports Base First_Steps", "begin"] |
|
346
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
10 |
*} |
|
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
11 |
(*>*) |
|
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
12 |
|
|
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
345
diff
changeset
|
13 |
|
| 414 | 14 |
chapter {* Advanced Isabelle\label{chp:advanced} *}
|
|
381
97518188ef0e
added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents:
380
diff
changeset
|
15 |
|
|
97518188ef0e
added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents:
380
diff
changeset
|
16 |
text {*
|
| 410 | 17 |
\begin{flushright}
|
18 |
{\em All things are difficult before they are easy.} \\[1ex]
|
|
19 |
proverb |
|
20 |
\end{flushright}
|
|
21 |
||
22 |
\medskip |
|
| 400 | 23 |
While terms, types and theorems are the most basic data structures in |
24 |
Isabelle, there are a number of layers built on top of them. Most of these |
|
| 408 | 25 |
layers are concerned with storing and manipulating data. Handling them |
| 487 | 26 |
properly is an essential skill for programming on the ML-level of Isabelle. |
27 |
The most basic layer are theories. They contain global data and |
|
| 408 | 28 |
can be seen as the ``long-term memory'' of Isabelle. There is usually only |
29 |
one theory active at each moment. Proof contexts and local theories, on the |
|
| 400 | 30 |
other hand, store local data for a task at hand. They act like the |
| 408 | 31 |
``short-term memory'' and there can be many of them that are active in |
32 |
parallel. |
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
*} |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
35 |
section {* Theories and Setups\label{sec:theories} *}
|
| 358 | 36 |
|
37 |
text {*
|
|
| 492 | 38 |
Theories, as said above, are the most basic layer of abstraction in |
39 |
Isabelle. They record information about definitions, syntax declarations, axioms, |
|
| 488 | 40 |
theorems and much more. For example, if a definition is made, it |
41 |
must be stored in a theory in order to be usable later on. Similar |
|
42 |
with proofs: once a proof is finished, the proved theorem needs to |
|
43 |
be stored in the theorem database of the theory in order to be |
|
| 491 | 44 |
usable. All relevant data of a theory can be queried with the |
45 |
Isabelle command \isacommand{print\_theory}.
|
|
| 358 | 46 |
|
47 |
\begin{isabelle}
|
|
48 |
\isacommand{print\_theory}\\
|
|
49 |
@{text "> names: Pure Code_Generator HOL \<dots>"}\\
|
|
50 |
@{text "> classes: Inf < type \<dots>"}\\
|
|
51 |
@{text "> default sort: type"}\\
|
|
52 |
@{text "> syntactic types: #prop \<dots>"}\\
|
|
53 |
@{text "> logical types: 'a \<times> 'b \<dots>"}\\
|
|
54 |
@{text "> type arities: * :: (random, random) random \<dots>"}\\
|
|
55 |
@{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
|
|
56 |
@{text "> abbreviations: \<dots>"}\\
|
|
57 |
@{text "> axioms: \<dots>"}\\
|
|
58 |
@{text "> oracles: \<dots>"}\\
|
|
59 |
@{text "> definitions: \<dots>"}\\
|
|
60 |
@{text "> theorems: \<dots>"}
|
|
61 |
\end{isabelle}
|
|
62 |
||
| 487 | 63 |
Functions acting on theories often end with the suffix @{text "_global"},
|
64 |
for example the function @{ML read_term_global in Syntax} in the structure
|
|
65 |
@{ML_struct Syntax}. The reason is to set them syntactically apart from
|
|
| 491 | 66 |
functions acting on contexts or local theories, which will be discussed in |
67 |
the next sections. There is a tendency amongst Isabelle developers to prefer |
|
| 487 | 68 |
``non-global'' operations, because they have some advantages, as we will also |
| 490 | 69 |
discuss later. However, some basic understanding of theories is still necessary |
70 |
for effective Isabelle programming. |
|
| 487 | 71 |
|
| 491 | 72 |
An important Isabelle command with theories is \isacommand{setup}. In the
|
| 492 | 73 |
previous chapters we used it already to make a theorem attribute known |
74 |
to Isabelle and to register a theorem under a name. What happens behind the |
|
| 490 | 75 |
scenes is that \isacommand{setup} expects a function of type @{ML_type
|
76 |
"theory -> theory"}: the input theory is the current theory and the output |
|
77 |
the theory where the attribute has been registered or the theorem has been |
|
78 |
stored. This is a fundamental principle in Isabelle. A similar situation |
|
| 491 | 79 |
arises with declaring a constant, which can be done on the ML-level with |
80 |
function @{ML_ind declare_const in Sign} from the structure @{ML_struct
|
|
| 490 | 81 |
Sign}. To see how \isacommand{setup} works, consider the following code:
|
82 |
||
| 348 | 83 |
*} |
84 |
||
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
85 |
ML{*let
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
86 |
val thy = @{theory}
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
87 |
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
88 |
in |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
89 |
Sign.declare_const @{context} bar_const thy
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
90 |
end*} |
| 348 | 91 |
|
92 |
text {*
|
|
| 488 | 93 |
If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
94 |
\isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} with the
|
| 491 | 95 |
intention of declaring a constant @{text "BAR"} having type @{typ nat}, then
|
| 488 | 96 |
indeed you obtain a theory as result. But if you query the |
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
97 |
constant on the Isabelle level using the command \isacommand{term}
|
| 348 | 98 |
|
99 |
\begin{isabelle}
|
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
100 |
\isacommand{term}~@{text BAR}\\
|
| 348 | 101 |
@{text "> \"BAR\" :: \"'a\""}
|
102 |
\end{isabelle}
|
|
103 |
||
| 495 | 104 |
you can see that you do \emph{not} obtain the expected constant of type @{typ nat}, but a free
|
| 484 | 105 |
variable (printed in blue) of polymorphic type. The problem is that the |
106 |
ML-expression above did not ``register'' the declaration with the current theory. |
|
107 |
This is what the command \isacommand{setup} is for. The constant is properly
|
|
108 |
declared with |
|
| 348 | 109 |
*} |
110 |
||
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
111 |
setup %gray {* fn thy =>
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
112 |
let |
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
113 |
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
114 |
val (_, thy') = Sign.declare_const @{context} bar_const thy
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
115 |
in |
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
116 |
thy' |
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
117 |
end *} |
| 348 | 118 |
|
119 |
text {*
|
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
120 |
where the declaration is actually applied to the current theory and |
| 348 | 121 |
|
122 |
\begin{isabelle}
|
|
123 |
\isacommand{term}~@{text [quotes] "BAR"}\\
|
|
124 |
@{text "> \"BAR\" :: \"nat\""}
|
|
125 |
\end{isabelle}
|
|
126 |
||
| 492 | 127 |
now returns a (black) constant with the type @{typ nat}, as expected.
|
| 348 | 128 |
|
| 491 | 129 |
In a sense, \isacommand{setup} can be seen as a transaction that
|
130 |
takes the current theory @{text thy}, applies an operation, and
|
|
131 |
produces a new current theory @{text thy'}. This means that we have
|
|
132 |
to be careful to apply operations always to the most current theory, |
|
133 |
not to a \emph{stale} one. Consider again the function inside the
|
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
134 |
\isacommand{setup}-command:
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
135 |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
136 |
\begin{isabelle}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
137 |
\begin{graybox}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
138 |
\isacommand{setup}~@{text "\<verbopen>"} @{ML
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
139 |
"fn thy => |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
140 |
let |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
141 |
val bar_const = ((@{binding \"BAR\"}, @{typ \"nat\"}), NoSyn)
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
142 |
val (_, thy') = Sign.declare_const @{context} bar_const thy
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
143 |
in |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
144 |
thy |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
145 |
end"}~@{text "\<verbclose>"}\isanewline
|
| 495 | 146 |
@{text "> ERROR: \"Stale theory encountered\""}
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
147 |
\end{graybox}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
148 |
\end{isabelle}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
149 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
150 |
This time we erroneously return the original theory @{text thy}, instead of
|
| 488 | 151 |
the modified one @{text thy'}. Such buggy code will always result into
|
152 |
a runtime error message about stale theories. |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
153 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
154 |
However, sometimes it does make sense to work with two theories at the same |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
155 |
time, especially in the context of parsing and typing. In the code below we |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
156 |
use in Line 3 the function @{ML_ind copy in Theory} from the structure
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
157 |
@{ML_struct Theory} for obtaining a new theory that contains the same
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
158 |
data, but is unrelated to the existing theory. |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
159 |
*} |
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
160 |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
161 |
setup %graylinenos {* fn thy =>
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
162 |
let |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
163 |
val tmp_thy = Theory.copy thy |
| 498 | 164 |
val foo_const = ((@{binding "FOO"}, @{typ "nat \<Rightarrow> nat"}), NoSyn)
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
165 |
val (_, tmp_thy') = Sign.declare_const @{context} foo_const tmp_thy
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
166 |
val trm1 = Syntax.read_term_global tmp_thy' "FOO baz" |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
167 |
val trm2 = Syntax.read_term_global thy "FOO baz" |
| 502 | 168 |
val _ = pwriteln |
169 |
(Pretty.str (@{make_string} trm1 ^ "\n" ^ @{make_string} trm2))
|
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
170 |
in |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
171 |
thy |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
172 |
end *} |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
173 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
174 |
text {*
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
175 |
That means we can make changes to the theory @{text tmp_thy} without
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
176 |
affecting the current theory @{text thy}. In this case we declare in @{text
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
177 |
"tmp_thy"} the constant @{text FOO} (Lines 4 and 5). The point of this code
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
178 |
is that we next, in Lines 6 and 7, parse a string to become a term (both |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
179 |
times the string is @{text [quotes] "FOO baz"}). But since we parse the string
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
180 |
once in the context of the theory @{text tmp_thy'} in which @{text FOO} is
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
181 |
declared to be a constant of type @{typ "nat \<Rightarrow>nat"} and once in the context
|
| 488 | 182 |
of @{text thy} where it is not, we obtain two different terms, namely
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
183 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
184 |
\begin{isabelle}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
185 |
\begin{graybox}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
186 |
@{text "> Const (\"Advanced.FOO\", \"nat \<Rightarrow> nat\") $ Free (\"baz\", \"nat\")"}\isanewline
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
187 |
@{text "> Free (\"FOO\", \"'a \<Rightarrow> 'b\") $ Free (\"baz\", \"'a\")"}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
188 |
\end{graybox}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
189 |
\end{isabelle}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
190 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
191 |
There are two reasons for parsing a term in a temporary theory. One is to |
| 488 | 192 |
obtain fully qualified names for constants and the other is appropriate type |
193 |
inference. This is relevant in situations where definitions are made later, |
|
194 |
but parsing and type inference has to take already proceed as if the definitions |
|
195 |
were already made. |
|
| 502 | 196 |
|
197 |
\begin{readmore}
|
|
198 |
Most of the functions about theories are implemented in |
|
199 |
@{ML_file "Pure/theory.ML"} and @{ML_file "Pure/global_theory.ML"}.
|
|
200 |
\end{readmore}
|
|
| 348 | 201 |
*} |
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
202 |
|
| 495 | 203 |
section {* Contexts *}
|
|
341
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
204 |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
205 |
text {*
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
206 |
Contexts are arguably more important than theories, even though they only |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
207 |
contain ``short-term memory data''. The reason is that a vast number of |
| 490 | 208 |
functions in Isabelle depend in one way or another on contexts. Even such |
209 |
mundane operations like printing out a term make essential use of contexts. |
|
| 494 | 210 |
For this consider the following contrived proof-snippet whose purpose is to |
| 490 | 211 |
fix two variables: |
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
212 |
*} |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
213 |
|
| 488 | 214 |
lemma "True" |
215 |
proof - |
|
| 491 | 216 |
txt_raw {*\mbox{}\\[-7mm]*}
|
| 492 | 217 |
ML_prf {* Variable.dest_fixes @{context} *}
|
| 491 | 218 |
txt_raw {*\mbox{}\\[-7mm]\mbox{}*}
|
| 492 | 219 |
fix x y |
| 491 | 220 |
txt_raw {*\mbox{}\\[-7mm]*}
|
| 488 | 221 |
ML_prf {* Variable.dest_fixes @{context} *}
|
| 491 | 222 |
txt_raw {*\mbox{}\\[-7mm] \ldots*}(*<*)oops(*>*)
|
| 488 | 223 |
|
| 490 | 224 |
text {*
|
| 496 | 225 |
The interesting point is that we injected ML-code before and after |
| 490 | 226 |
the variables are fixed. For this remember that ML-code inside a proof |
| 494 | 227 |
needs to be enclosed inside \isacommand{ML\_prf}~@{text "\<verbopen> \<dots> \<verbclose>"},
|
| 490 | 228 |
not \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}. The function
|
229 |
@{ML_ind dest_fixes in Variable} from the structure @{ML_struct Variable} takes
|
|
230 |
a context and returns all its currently fixed variable (names). That |
|
231 |
means a context has a dataslot containing information about fixed variables. |
|
232 |
The ML-antiquotation @{text "@{context}"} points to the context that is
|
|
233 |
active at that point of the theory. Consequently, in the first call to |
|
234 |
@{ML dest_fixes in Variable} this dataslot is empty; in the second it is
|
|
235 |
filled with @{text x} and @{text y}. What is interesting is that contexts
|
|
| 494 | 236 |
can be stacked. For this consider the following proof fragment: |
| 490 | 237 |
*} |
238 |
||
239 |
lemma "True" |
|
240 |
proof - |
|
241 |
fix x y |
|
242 |
{ fix z w
|
|
| 491 | 243 |
txt_raw {*\mbox{}\\[-7mm]*}
|
244 |
ML_prf {* Variable.dest_fixes @{context} *}
|
|
| 492 | 245 |
txt_raw {*\mbox{}\\[-7mm]\mbox{}*}
|
246 |
} |
|
| 491 | 247 |
txt_raw {*\mbox{}\\[-7mm]*}
|
| 490 | 248 |
ML_prf {* Variable.dest_fixes @{context} *}
|
| 491 | 249 |
txt_raw {*\mbox{}\\[-7mm] \ldots*}(*<*)oops(*>*)
|
250 |
||
251 |
text {*
|
|
| 495 | 252 |
Here the first time we call @{ML dest_fixes in Variable} we have four fixes variables;
|
253 |
the second time we get only the fixes variables @{text x} and @{text y} as answer, since
|
|
254 |
@{text z} and @{text w} are not anymore in the scope of the context.
|
|
255 |
This means the curly-braces act on the Isabelle level as opening and closing statements |
|
| 496 | 256 |
for a context. The above proof fragment corresponds roughly to the following ML-code |
| 491 | 257 |
*} |
| 488 | 258 |
|
| 492 | 259 |
ML{*val ctxt0 = @{context};
|
260 |
val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0; |
|
261 |
val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1*} |
|
262 |
||
263 |
text {*
|
|
| 494 | 264 |
where the function @{ML_ind add_fixes in Variable} fixes a list of variables
|
265 |
specified by strings. Let us come back to the point about printing terms. Consider |
|
| 495 | 266 |
printing out the term \mbox{@{text "(x, y, z, w)"}} using our function @{ML_ind pretty_term}.
|
| 496 | 267 |
This function takes a term and a context as argument. Notice how the printing |
| 498 | 268 |
of the term changes according to which context is used. |
| 492 | 269 |
|
| 495 | 270 |
\begin{isabelle}
|
271 |
\begin{graybox}
|
|
272 |
@{ML "let
|
|
273 |
val trm = @{term \"(x, y, z, w)\"}
|
|
| 493 | 274 |
in |
275 |
pwriteln (Pretty.chunks |
|
276 |
[ pretty_term ctxt0 trm, |
|
277 |
pretty_term ctxt1 trm, |
|
278 |
pretty_term ctxt2 trm ]) |
|
| 495 | 279 |
end"}\\ |
280 |
\setlength{\fboxsep}{0mm}
|
|
281 |
@{text ">"}~@{text "("}\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
|
|
282 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
|
|
283 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
|
|
284 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}\\
|
|
285 |
@{text ">"}~@{text "("}\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
|
|
286 |
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
|
|
287 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
|
|
288 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}\\
|
|
289 |
@{text ">"}~@{text "("}\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text x}}}@{text ","}~%
|
|
290 |
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text y}}}@{text ","}~%
|
|
291 |
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text z}}}@{text ","}~%
|
|
292 |
\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{@{text w}}}@{text ")"}
|
|
293 |
\end{graybox}
|
|
294 |
\end{isabelle}
|
|
295 |
||
296 |
||
297 |
The term we are printing is in all three cases the same---a tuple containing |
|
| 496 | 298 |
four free variables. As you can see, however, in case of @{ML "ctxt0"} all
|
| 495 | 299 |
variables are highlighted indicating a problem, while in case of @{ML
|
300 |
"ctxt1"} @{text x} and @{text y} are printed as normal (blue) free
|
|
301 |
variables, but not @{text z} and @{text w}. In the last case all variables
|
|
302 |
are printed as expected. The point of this example is that the context |
|
303 |
contains the information which variables are fixed, and designates all other |
|
| 497 | 304 |
free variables as being alien or faulty. Therefore the highlighting. |
305 |
While this seems like a minor detail, the concept of making the context aware |
|
306 |
of fixed variables is actually quite useful. For example it prevents us from |
|
| 501 | 307 |
fixing a variable twice |
| 495 | 308 |
|
309 |
@{ML_response_fake [gray, display]
|
|
310 |
"@{context}
|
|
311 |
|> Variable.add_fixes [\"x\", \"x\"]" |
|
312 |
"ERROR: Duplicate fixed variable(s): \"x\""} |
|
313 |
||
| 501 | 314 |
More importantly it also allows us to easily create fresh names for |
315 |
fixed variables. For this you have to use the function @{ML_ind
|
|
316 |
variant_fixes in Variable} from the structure @{ML_struct Variable}.
|
|
317 |
||
318 |
@{ML_response_fake [gray, display]
|
|
319 |
"@{context}
|
|
320 |
|> Variable.variant_fixes [\"y\", \"y\", \"z\"]" |
|
321 |
"([\"y\", \"ya\", \"z\"], ...)"} |
|
322 |
||
323 |
Now a fresh variant for the second occurence of @{text y} is created
|
|
| 502 | 324 |
avoiding any clash. In this way we can also create fresh free variables |
| 501 | 325 |
that avoid any clashes with fixed variables. In Line~3 below we fix |
326 |
the variable @{text x} in the context @{text ctxt1}. Next we want to
|
|
327 |
create two fresh variables of type @{typ nat} as variants of the
|
|
328 |
string @{text [quotes] "x"} (Lines 6 and 7).
|
|
| 495 | 329 |
|
330 |
@{ML_response_fake [display, gray, linenos]
|
|
331 |
"let |
|
332 |
val ctxt0 = @{context}
|
|
333 |
val (_, ctxt1) = Variable.add_fixes [\"x\"] ctxt0 |
|
334 |
val frees = replicate 2 (\"x\", @{typ nat})
|
|
335 |
in |
|
336 |
(Variable.variant_frees ctxt0 [] frees, |
|
337 |
Variable.variant_frees ctxt1 [] frees) |
|
338 |
end" |
|
339 |
"([(\"x\", \"nat\"), (\"xa\", \"nat\")], |
|
340 |
[(\"xa\", \"nat\"), (\"xb\", \"nat\")])"} |
|
341 |
||
| 496 | 342 |
As you can see, if we create the fresh variables with the context @{text ctxt0},
|
| 495 | 343 |
then we obtain @{text "x"} and @{text "xa"}; but in @{text ctxt1} we obtain @{text "xa"}
|
344 |
and @{text "xb"} avoiding @{text x}, which was fixed in this context.
|
|
345 |
||
| 496 | 346 |
Often one has the problem that given some terms, create fresh variables |
347 |
avoiding any variable occurring in those terms. For this you can use the |
|
| 495 | 348 |
function @{ML_ind declare_term in Variable} from the structure @{ML_struct Variable}.
|
349 |
||
350 |
@{ML_response_fake [gray, display]
|
|
351 |
"let |
|
352 |
val ctxt1 = Variable.declare_term @{term \"(x, xa)\"} @{context}
|
|
353 |
val frees = replicate 2 (\"x\", @{typ nat})
|
|
354 |
in |
|
355 |
Variable.variant_frees ctxt1 [] frees |
|
356 |
end" |
|
357 |
"[(\"xb\", \"nat\"), (\"xc\", \"nat\")]"} |
|
358 |
||
| 496 | 359 |
The result is @{text xb} and @{text xc} for the names of the fresh
|
| 498 | 360 |
variables, since @{text x} and @{text xa} occur in the term we declared.
|
361 |
Note that @{ML_ind declare_term in Variable} does not fix the
|
|
362 |
variables; it just makes them ``known'' to the context. You can see |
|
363 |
that if you print out a declared term. |
|
364 |
||
365 |
\begin{isabelle}
|
|
366 |
\begin{graybox}
|
|
367 |
@{ML "let
|
|
368 |
val trm = @{term \"P x y z\"}
|
|
369 |
val ctxt1 = Variable.declare_term trm @{context}
|
|
370 |
in |
|
371 |
pwriteln (pretty_term ctxt1 trm) |
|
372 |
end"}\\ |
|
373 |
\setlength{\fboxsep}{0mm}
|
|
374 |
@{text ">"}~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text P}}}~%
|
|
375 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text x}}}~%
|
|
376 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text y}}}~%
|
|
377 |
\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text z}}}
|
|
378 |
\end{graybox}
|
|
379 |
\end{isabelle}
|
|
380 |
||
381 |
All variables are highligted, indicating that they are not |
|
382 |
fixed. However, declaring a term is helpful when parsing terms using |
|
383 |
the function @{ML_ind read_term in Syntax} from the structure
|
|
384 |
@{ML_struct Syntax}. Consider the following code:
|
|
| 495 | 385 |
|
386 |
@{ML_response_fake [gray, display]
|
|
387 |
"let |
|
388 |
val ctxt0 = @{context}
|
|
389 |
val ctxt1 = Variable.declare_term @{term \"x::nat\"} ctxt0
|
|
390 |
in |
|
391 |
(Syntax.read_term ctxt0 \"x\", |
|
392 |
Syntax.read_term ctxt1 \"x\") |
|
393 |
end" |
|
394 |
"(Free (\"x\", \"'a\"), Free (\"x\", \"nat\"))"} |
|
395 |
||
| 496 | 396 |
Parsing the string in the context @{text "ctxt0"} results in a free variable
|
397 |
with a default polymorphic type, but in case of @{text "ctxt1"} we obtain a
|
|
| 495 | 398 |
free variable of type @{typ nat} as previously declared. Which
|
| 496 | 399 |
type the parsed term receives depends upon the last declaration that |
400 |
is made, as the next example illustrates. |
|
| 495 | 401 |
|
402 |
@{ML_response_fake [gray, display]
|
|
403 |
"let |
|
404 |
val ctxt1 = Variable.declare_term @{term \"x::nat\"} @{context}
|
|
405 |
val ctxt2 = Variable.declare_term @{term \"x::int\"} ctxt1
|
|
406 |
in |
|
407 |
(Syntax.read_term ctxt1 \"x\", |
|
408 |
Syntax.read_term ctxt2 \"x\") |
|
409 |
end" |
|
410 |
"(Free (\"x\", \"nat\"), Free (\"x\", \"int\"))"} |
|
411 |
||
| 499 | 412 |
The most useful feature of contexts is that one can export, or transfer, |
413 |
terms and theorems between them. We show this first for terms. |
|
| 497 | 414 |
|
415 |
\begin{isabelle}
|
|
416 |
\begin{graybox}
|
|
417 |
\begin{linenos}
|
|
418 |
@{ML "let
|
|
419 |
val ctxt0 = @{context}
|
|
420 |
val (_, ctxt1) = Variable.add_fixes [\"x\", \"y\", \"z\"] ctxt0 |
|
421 |
val foo_trm = @{term \"P x y z\"}
|
|
422 |
in |
|
423 |
singleton (Variable.export_terms ctxt1 ctxt0) foo_trm |
|
424 |
|> pretty_term ctxt0 |
|
425 |
|> pwriteln |
|
426 |
end"} |
|
427 |
\end{linenos}
|
|
428 |
\setlength{\fboxsep}{0mm}
|
|
429 |
@{text ">"}~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{@{text P}}}~%
|
|
430 |
@{text "?x ?y ?z"}
|
|
431 |
\end{graybox}
|
|
432 |
\end{isabelle}
|
|
433 |
||
| 498 | 434 |
In Line 3 we fix the variables @{term x}, @{term y} and @{term z} in
|
435 |
context @{text ctxt1}. The function @{ML_ind export_terms in
|
|
| 500 | 436 |
Variable} from the structure @{ML_struct Variable} can be used to transfer
|
| 498 | 437 |
terms between contexts. Transferring means to turn all (free) |
438 |
variables that are fixed in one context, but not in the other, into |
|
439 |
schematic variables. In our example, we are transferring the term |
|
440 |
@{text "P x y z"} from context @{text "ctxt1"} to @{text "ctxt0"},
|
|
441 |
which means @{term x}, @{term y} and @{term z} become schematic
|
|
| 500 | 442 |
variables (as can be seen by the leading question marks in the result). |
443 |
Note that the variable @{text P} stays a free variable, since it not fixed in
|
|
| 498 | 444 |
@{text ctxt1}; it is even highlighed, because @{text "ctxt0"} does
|
445 |
not know about it. Note also that in Line 6 we had to use the |
|
446 |
function @{ML_ind singleton}, because the function @{ML_ind
|
|
447 |
export_terms in Variable} normally works over lists of terms. |
|
448 |
||
449 |
The case of transferring theorems is even more useful. The reason is |
|
450 |
that the generalisation of fixed variables to schematic variables is |
|
| 499 | 451 |
not trivial if done manually. For illustration purposes we use in the |
452 |
following code the function @{ML_ind make_thm in Skip_Proof} from the
|
|
| 500 | 453 |
structure @{ML_struct Skip_Proof}. This function will turn an arbitray
|
454 |
term, in our case @{term "P x y z x y z"}, into a theorem (disregarding
|
|
455 |
whether it is actually provable). |
|
| 498 | 456 |
|
457 |
@{ML_response_fake [display, gray]
|
|
458 |
"let |
|
459 |
val thy = @{theory}
|
|
460 |
val ctxt0 = @{context}
|
|
| 499 | 461 |
val (_, ctxt1) = Variable.add_fixes [\"P\", \"x\", \"y\", \"z\"] ctxt0 |
| 498 | 462 |
val foo_thm = Skip_Proof.make_thm thy @{prop \"P x y z x y z\"}
|
463 |
in |
|
464 |
singleton (Proof_Context.export ctxt1 ctxt0) foo_thm |
|
465 |
end" |
|
| 499 | 466 |
"?P ?x ?y ?z ?x ?y ?z"} |
467 |
||
| 502 | 468 |
Since we fixed all variables in @{text ctxt1}, in the exported
|
469 |
result all of them are schematic. The great point of contexts is |
|
470 |
that exporting from one to another is not just restricted to |
|
471 |
variables, but also works with assumptions. For this we can use the |
|
472 |
function @{ML_ind export in Assumption} from the structure
|
|
| 500 | 473 |
@{ML_struct Assumption}. Consider the following code.
|
474 |
||
475 |
@{ML_response_fake [display, gray, linenos]
|
|
476 |
"let |
|
477 |
val ctxt0 = @{context}
|
|
478 |
val ([eq], ctxt1) = Assumption.add_assumes [@{cprop \"x \<equiv> y\"}] ctxt0
|
|
479 |
val eq' = Thm.symmetric eq |
|
480 |
in |
|
481 |
Assumption.export false ctxt1 ctxt0 eq' |
|
482 |
end" |
|
483 |
"x \<equiv> y \<Longrightarrow> y \<equiv> x"} |
|
484 |
||
485 |
The function @{ML_ind add_assumes in Assumption} from the structure
|
|
486 |
@{ML_struct Assumption} adds the assumption \mbox{@{text "x \<equiv> y"}}
|
|
487 |
to the context @{text ctxt1} (Line 3). This function expects a list
|
|
488 |
of @{ML_type cterm}s and returns them as theorems, together with the
|
|
489 |
new context in which they are ``assumed''. In Line 4 we use the |
|
490 |
function @{ML_ind symmetric in Thm} from the structure @{ML_struct
|
|
491 |
Thm} in order to obtain the symmetric version of the assumed |
|
492 |
meta-equality. Now exporting the theorem @{text "eq'"} from @{text
|
|
493 |
ctxt1} to @{text ctxt0} means @{term "y \<equiv> x"} will be prefixed with
|
|
494 |
the assumed theorem. The boolean flag in @{ML_ind export in
|
|
495 |
Assumption} indicates whether the assumptions should be marked with |
|
496 |
the goal marker (see Section~\ref{sec:basictactics}). In normal
|
|
497 |
circumstances this is not necessary and so should be set to @{ML
|
|
498 |
false}. The result of the export is then the theorem \mbox{@{term
|
|
499 |
"x \<equiv> y \<Longrightarrow> y \<equiv> x"}}. As can be seen this is an easy way for obtaing |
|
500 |
simple theorems. We will explain this in more detail in |
|
501 |
Section~\ref{sec:structured}.
|
|
502 |
||
503 |
The function @{ML_ind export in Proof_Context} from the structure
|
|
504 |
@{ML_struct Proof_Context} combines both export functions from
|
|
505 |
@{ML_struct Variable} and @{ML_struct Assumption}. This can be seen
|
|
506 |
in the following example. |
|
507 |
||
508 |
@{ML_response_fake [display, gray]
|
|
509 |
"let |
|
510 |
val ctxt0 = @{context}
|
|
511 |
val ((fvs, [eq]), ctxt1) = ctxt0 |
|
512 |
|> Variable.add_fixes [\"x\", \"y\"] |
|
513 |
||>> Assumption.add_assumes [@{cprop \"x \<equiv> y\"}]
|
|
514 |
val eq' = Thm.symmetric eq |
|
515 |
in |
|
516 |
Proof_Context.export ctxt1 ctxt0 [eq'] |
|
517 |
end" |
|
518 |
"[?x \<equiv> ?y \<Longrightarrow> ?y \<equiv> ?x]"} |
|
| 495 | 519 |
*} |
520 |
||
| 496 | 521 |
|
| 493 | 522 |
|
| 495 | 523 |
text {*
|
| 493 | 524 |
|
525 |
*} |
|
526 |
||
| 492 | 527 |
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
528 |
(* |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
529 |
ML{*Proof_Context.debug := true*}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
530 |
ML{*Proof_Context.verbose := true*}
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
531 |
*) |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
532 |
|
| 487 | 533 |
(* |
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
534 |
lemma "True" |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
535 |
proof - |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
536 |
{ -- "\<And>x. _"
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
537 |
fix x |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
538 |
have "B x" sorry |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
539 |
thm this |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
540 |
} |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
541 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
542 |
thm this |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
543 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
544 |
{ -- "A \<Longrightarrow> _"
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
545 |
assume A |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
546 |
have B sorry |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
547 |
thm this |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
548 |
} |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
549 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
550 |
thm this |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
551 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
552 |
{ -- "\<And>x. x = _ \<Longrightarrow> _"
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
553 |
def x \<equiv> a |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
554 |
have "B x" sorry |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
555 |
} |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
556 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
557 |
thm this |
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
558 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
559 |
oops |
| 487 | 560 |
*) |
| 413 | 561 |
|
| 502 | 562 |
section {* Local Theories and Local Setups\label{sec:local} (TBD) *}
|
|
341
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
563 |
|
| 394 | 564 |
text {*
|
| 400 | 565 |
In contrast to an ordinary theory, which simply consists of a type |
566 |
signature, as well as tables for constants, axioms and theorems, a local |
|
567 |
theory contains additional context information, such as locally fixed |
|
568 |
variables and local assumptions that may be used by the package. The type |
|
569 |
@{ML_type local_theory} is identical to the type of \emph{proof contexts}
|
|
570 |
@{ML_type "Proof.context"}, although not every proof context constitutes a
|
|
571 |
valid local theory. |
|
572 |
||
573 |
@{ML "Context.>> o Context.map_theory"}
|
|
| 394 | 574 |
@{ML_ind "Local_Theory.declaration"}
|
|
486
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
575 |
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
576 |
A similar command is \isacommand{local\_setup}, which expects a function
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
577 |
of type @{ML_type "local_theory -> local_theory"}. Later on we will also
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
578 |
use the commands \isacommand{method\_setup} for installing methods in the
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
579 |
current theory and \isacommand{simproc\_setup} for adding new simprocs to
|
|
45cfd2ece7bd
a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents:
485
diff
changeset
|
580 |
the current simpset. |
| 394 | 581 |
*} |
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
582 |
|
|
336
a12bb28fe2bd
polished on the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents:
335
diff
changeset
|
583 |
|
| 394 | 584 |
section {* Morphisms (TBD) *}
|
585 |
||
586 |
text {*
|
|
587 |
Morphisms are arbitrary transformations over terms, types, theorems and bindings. |
|
588 |
They can be constructed using the function @{ML_ind morphism in Morphism},
|
|
589 |
which expects a record with functions of type |
|
590 |
||
591 |
\begin{isabelle}
|
|
592 |
\begin{tabular}{rl}
|
|
593 |
@{text "binding:"} & @{text "binding -> binding"}\\
|
|
594 |
@{text "typ:"} & @{text "typ -> typ"}\\
|
|
595 |
@{text "term:"} & @{text "term -> term"}\\
|
|
596 |
@{text "fact:"} & @{text "thm list -> thm list"}
|
|
597 |
\end{tabular}
|
|
598 |
\end{isabelle}
|
|
599 |
||
600 |
The simplest morphism is the @{ML_ind identity in Morphism}-morphism defined as
|
|
601 |
*} |
|
602 |
||
| 481 | 603 |
ML{*val identity = Morphism.morphism {binding = [], typ = [], term = [], fact = []}*}
|
| 394 | 604 |
|
605 |
text {*
|
|
606 |
Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
|
|
607 |
*} |
|
608 |
||
609 |
ML{*fun trm_phi (Free (x, T)) = Var ((x, 0), T)
|
|
610 |
| trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t) |
|
611 |
| trm_phi (t $ s) = (trm_phi t) $ (trm_phi s) |
|
612 |
| trm_phi t = t*} |
|
613 |
||
614 |
ML{*val phi = Morphism.term_morphism trm_phi*}
|
|
615 |
||
616 |
ML{*Morphism.term phi @{term "P x y"}*}
|
|
617 |
||
618 |
text {*
|
|
619 |
@{ML_ind term_morphism in Morphism}
|
|
620 |
||
621 |
@{ML_ind term in Morphism},
|
|
622 |
@{ML_ind thm in Morphism}
|
|
623 |
||
624 |
\begin{readmore}
|
|
625 |
Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
|
|
626 |
\end{readmore}
|
|
627 |
*} |
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
628 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
629 |
section {* Misc (TBD) *}
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
630 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
631 |
ML {*Datatype.get_info @{theory} "List.list"*}
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
632 |
|
|
319
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
633 |
text {*
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
634 |
FIXME: association lists: |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
635 |
@{ML_file "Pure/General/alist.ML"}
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
636 |
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
637 |
FIXME: calling the ML-compiler |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
638 |
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
639 |
*} |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
640 |
|
| 414 | 641 |
section {* What Is In an Isabelle Name? (TBD) *}
|
642 |
||
643 |
text {*
|
|
644 |
On the ML-level of Isabelle, you often have to work with qualified names. |
|
645 |
These are strings with some additional information, such as positional |
|
646 |
information and qualifiers. Such qualified names can be generated with the |
|
647 |
antiquotation @{text "@{binding \<dots>}"}. For example
|
|
648 |
||
649 |
@{ML_response [display,gray]
|
|
650 |
"@{binding \"name\"}"
|
|
651 |
"name"} |
|
652 |
||
653 |
An example where a qualified name is needed is the function |
|
654 |
@{ML_ind define in Local_Theory}. This function is used below to define
|
|
655 |
the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
|
|
656 |
*} |
|
657 |
||
658 |
local_setup %gray {*
|
|
659 |
Local_Theory.define ((@{binding "TrueConj"}, NoSyn),
|
|
660 |
(Attrib.empty_binding, @{term "True \<and> True"})) #> snd *}
|
|
661 |
||
662 |
text {*
|
|
663 |
Now querying the definition you obtain: |
|
664 |
||
665 |
\begin{isabelle}
|
|
666 |
\isacommand{thm}~@{text "TrueConj_def"}\\
|
|
667 |
@{text "> "}~@{thm TrueConj_def}
|
|
668 |
\end{isabelle}
|
|
669 |
||
670 |
\begin{readmore}
|
|
671 |
The basic operations on bindings are implemented in |
|
672 |
@{ML_file "Pure/General/binding.ML"}.
|
|
673 |
\end{readmore}
|
|
674 |
||
675 |
\footnote{\bf FIXME give a better example why bindings are important}
|
|
676 |
\footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
|
|
677 |
why @{ML snd} is needed.}
|
|
678 |
\footnote{\bf FIXME: There should probably a separate section on binding, long-names
|
|
679 |
and sign.} |
|
680 |
||
681 |
*} |
|
682 |
||
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
683 |
|
| 360 | 684 |
ML {* Sign.intern_type @{theory} "list" *}
|
685 |
ML {* Sign.intern_const @{theory} "prod_fun" *}
|
|
686 |
||
| 414 | 687 |
text {*
|
688 |
\footnote{\bf FIXME: Explain the following better; maybe put in a separate
|
|
689 |
section and link with the comment in the antiquotation section.} |
|
690 |
||
691 |
Occasionally you have to calculate what the ``base'' name of a given |
|
| 462 | 692 |
constant is. For this you can use the function @{ML_ind Long_Name.base_name}. For example:
|
| 414 | 693 |
|
| 462 | 694 |
@{ML_response [display,gray] "Long_Name.base_name \"List.list.Nil\"" "\"Nil\""}
|
| 414 | 695 |
|
696 |
\begin{readmore}
|
|
697 |
Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
|
|
698 |
functions about signatures in @{ML_file "Pure/sign.ML"}.
|
|
699 |
\end{readmore}
|
|
700 |
*} |
|
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
701 |
|
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
702 |
text {*
|
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
703 |
@{ML_ind "Binding.name_of"} returns the string without markup
|
| 394 | 704 |
|
705 |
@{ML_ind "Binding.conceal"}
|
|
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
706 |
*} |
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
707 |
|
|
388
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
708 |
section {* Concurrency (TBD) *}
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
709 |
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
710 |
text {*
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
711 |
@{ML_ind prove_future in Goal}
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
712 |
@{ML_ind future_result in Goal}
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
713 |
@{ML_ind fork_pri in Future}
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
714 |
*} |
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
715 |
|
| 396 | 716 |
section {* Parse and Print Translations (TBD) *}
|
717 |
||
|
349
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
718 |
section {* Summary *}
|
|
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
719 |
|
|
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
720 |
text {*
|
|
395
2c392f61f400
spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents:
394
diff
changeset
|
721 |
TBD |
|
349
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
722 |
*} |
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
723 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
724 |
end |