| author | Christian Urban <urbanc@in.tum.de> |
| Sat, 05 Nov 2011 18:44:28 +0000 | |
| changeset 485 | f3536f0b47a9 |
| parent 484 | 490fe9483c0d |
| child 486 | 45cfd2ece7bd |
| 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 |
26 |
properly is an essential skill for programming on the ML-level of Isabelle |
|
27 |
programming. The most basic layer are theories. They contain global data and |
|
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 |
|
| 358 | 35 |
section {* Theories\label{sec:theories} (TBD) *}
|
36 |
||
37 |
text {*
|
|
|
401
36d61044f9bf
updated to new Isabelle and clarified Skip_Proof
Christian Urban <urbanc@in.tum.de>
parents:
400
diff
changeset
|
38 |
Theories, as said above, are the most basic layer in Isabelle. They contain |
| 400 | 39 |
definitions, syntax declarations, axioms, proofs etc. If a definition is |
40 |
stated, it must be stored in a theory in order to be usable later |
|
41 |
on. Similar with proofs: once a proof is finished, the proved theorem needs |
|
42 |
to be stored in the theorem database of the theory in order to be |
|
43 |
usable. All relevant data of a theory can be queried as follows. |
|
| 358 | 44 |
|
45 |
\begin{isabelle}
|
|
46 |
\isacommand{print\_theory}\\
|
|
47 |
@{text "> names: Pure Code_Generator HOL \<dots>"}\\
|
|
48 |
@{text "> classes: Inf < type \<dots>"}\\
|
|
49 |
@{text "> default sort: type"}\\
|
|
50 |
@{text "> syntactic types: #prop \<dots>"}\\
|
|
51 |
@{text "> logical types: 'a \<times> 'b \<dots>"}\\
|
|
52 |
@{text "> type arities: * :: (random, random) random \<dots>"}\\
|
|
53 |
@{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
|
|
54 |
@{text "> abbreviations: \<dots>"}\\
|
|
55 |
@{text "> axioms: \<dots>"}\\
|
|
56 |
@{text "> oracles: \<dots>"}\\
|
|
57 |
@{text "> definitions: \<dots>"}\\
|
|
58 |
@{text "> theorems: \<dots>"}
|
|
59 |
\end{isabelle}
|
|
60 |
||
| 400 | 61 |
|
62 |
||
| 358 | 63 |
\begin{center}
|
64 |
\begin{tikzpicture}
|
|
65 |
\node[top color=white, bottom color=gray!30, draw=black!100, drop shadow] {A};
|
|
66 |
\end{tikzpicture}
|
|
67 |
\end{center}
|
|
68 |
||
|
392
47e5b71c7f6c
modified the typ-pretty-printer and added parser exercise
Christian Urban <urbanc@in.tum.de>
parents:
389
diff
changeset
|
69 |
\footnote{\bf FIXME: list append in merge operations can cause
|
|
47e5b71c7f6c
modified the typ-pretty-printer and added parser exercise
Christian Urban <urbanc@in.tum.de>
parents:
389
diff
changeset
|
70 |
exponential blowups.} |
| 358 | 71 |
*} |
72 |
||
| 348 | 73 |
section {* Setups (TBD) *}
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
|
| 348 | 75 |
text {*
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
76 |
In the previous section we used \isacommand{setup} in order, for example,
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
77 |
to make a theorem attribute known to Isabelle or register a theorem under |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
78 |
a name. What happens behind the scenes is that \isacommand{setup} expects a
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
79 |
function of type @{ML_type "theory -> theory"}: the input theory is the current
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
80 |
theory and the output the theory where the theory attribute or theorem has been |
| 484 | 81 |
stored. |
| 361 | 82 |
|
| 484 | 83 |
This is a fundamental principle in Isabelle. A similar situation arises |
84 |
with declaring constants. The function that declares a |
|
85 |
constant on the ML-level is @{ML_ind declare_const in Sign}.
|
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
86 |
However, note that if you simply write\footnote{Recall that ML-code needs to be
|
| 348 | 87 |
enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.}
|
88 |
*} |
|
89 |
||
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
90 |
ML{*let
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
91 |
val thy = @{theory}
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
92 |
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
93 |
in |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
94 |
Sign.declare_const @{context} bar_const thy
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
95 |
end*} |
| 348 | 96 |
|
97 |
text {*
|
|
| 484 | 98 |
with the intention of declaring the constant @{text "BAR"} with type @{typ nat}
|
99 |
and run the code, then indeed you obtain a theory as result. But if you |
|
| 348 | 100 |
query the constant on the Isabelle level using the command \isacommand{term}
|
101 |
||
102 |
\begin{isabelle}
|
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
103 |
\isacommand{term}~@{text BAR}\\
|
| 348 | 104 |
@{text "> \"BAR\" :: \"'a\""}
|
105 |
\end{isabelle}
|
|
106 |
||
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
107 |
you can see that you do \emph{not} obtain a constant of type @{typ nat}, but a free
|
| 484 | 108 |
variable (printed in blue) of polymorphic type. The problem is that the |
109 |
ML-expression above did not ``register'' the declaration with the current theory. |
|
110 |
This is what the command \isacommand{setup} is for. The constant is properly
|
|
111 |
declared with |
|
| 348 | 112 |
*} |
113 |
||
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
114 |
setup %gray {* let
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
115 |
val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
116 |
in |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
117 |
Sign.declare_const @{context} bar_const #> snd
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
118 |
end *} |
| 348 | 119 |
|
120 |
text {*
|
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
121 |
where the declaration is actually applied to the theory and |
| 348 | 122 |
|
123 |
\begin{isabelle}
|
|
124 |
\isacommand{term}~@{text [quotes] "BAR"}\\
|
|
125 |
@{text "> \"BAR\" :: \"nat\""}
|
|
126 |
\end{isabelle}
|
|
127 |
||
| 484 | 128 |
returns a (black) constant with the type @{typ nat}, as expected.
|
| 348 | 129 |
|
|
485
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
130 |
In a sense, \isacommand{setup} can be seen as a transaction that takes the
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
131 |
current theory, applies an operation, and produces a new current theory. This |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
132 |
means that we have to be careful to apply operations always to the current |
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
133 |
theory, not to a \emph{stale} one. The code below produces
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
134 |
|
|
f3536f0b47a9
added section about testboard
Christian Urban <urbanc@in.tum.de>
parents:
484
diff
changeset
|
135 |
|
| 348 | 136 |
A similar command is \isacommand{local\_setup}, which expects a function
|
137 |
of type @{ML_type "local_theory -> local_theory"}. Later on we will also
|
|
138 |
use the commands \isacommand{method\_setup} for installing methods in the
|
|
139 |
current theory and \isacommand{simproc\_setup} for adding new simprocs to
|
|
140 |
the current simpset. |
|
141 |
*} |
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
142 |
|
|
341
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
143 |
section {* Contexts (TBD) *}
|
|
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
144 |
|
|
475
25371f74c768
updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents:
471
diff
changeset
|
145 |
ML{*Proof_Context.debug*}
|
|
25371f74c768
updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents:
471
diff
changeset
|
146 |
ML{*Proof_Context.verbose*}
|
| 413 | 147 |
|
148 |
||
|
341
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
149 |
section {* Local Theories (TBD) *}
|
|
62dea749d5ed
more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents:
340
diff
changeset
|
150 |
|
| 394 | 151 |
text {*
|
| 400 | 152 |
In contrast to an ordinary theory, which simply consists of a type |
153 |
signature, as well as tables for constants, axioms and theorems, a local |
|
154 |
theory contains additional context information, such as locally fixed |
|
155 |
variables and local assumptions that may be used by the package. The type |
|
156 |
@{ML_type local_theory} is identical to the type of \emph{proof contexts}
|
|
157 |
@{ML_type "Proof.context"}, although not every proof context constitutes a
|
|
158 |
valid local theory. |
|
159 |
||
160 |
@{ML "Context.>> o Context.map_theory"}
|
|
| 394 | 161 |
@{ML_ind "Local_Theory.declaration"}
|
162 |
*} |
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
163 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
164 |
(* |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
165 |
setup {*
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)]
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
*} |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
lemma "bar = (1::nat)" |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
oops |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
171 |
setup {*
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)]
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
#> PureThy.add_defs false [((@{binding "foo_def"},
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])]
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
#> snd |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
*} |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
*) |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
(* |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
lemma "foo = (1::nat)" |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
apply(simp add: foo_def) |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
done |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
thm foo_def |
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
*) |
|
336
a12bb28fe2bd
polished on the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents:
335
diff
changeset
|
185 |
|
| 394 | 186 |
section {* Morphisms (TBD) *}
|
187 |
||
188 |
text {*
|
|
189 |
Morphisms are arbitrary transformations over terms, types, theorems and bindings. |
|
190 |
They can be constructed using the function @{ML_ind morphism in Morphism},
|
|
191 |
which expects a record with functions of type |
|
192 |
||
193 |
\begin{isabelle}
|
|
194 |
\begin{tabular}{rl}
|
|
195 |
@{text "binding:"} & @{text "binding -> binding"}\\
|
|
196 |
@{text "typ:"} & @{text "typ -> typ"}\\
|
|
197 |
@{text "term:"} & @{text "term -> term"}\\
|
|
198 |
@{text "fact:"} & @{text "thm list -> thm list"}
|
|
199 |
\end{tabular}
|
|
200 |
\end{isabelle}
|
|
201 |
||
202 |
The simplest morphism is the @{ML_ind identity in Morphism}-morphism defined as
|
|
203 |
*} |
|
204 |
||
| 481 | 205 |
ML{*val identity = Morphism.morphism {binding = [], typ = [], term = [], fact = []}*}
|
| 394 | 206 |
|
207 |
text {*
|
|
208 |
Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
|
|
209 |
*} |
|
210 |
||
211 |
ML{*fun trm_phi (Free (x, T)) = Var ((x, 0), T)
|
|
212 |
| trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t) |
|
213 |
| trm_phi (t $ s) = (trm_phi t) $ (trm_phi s) |
|
214 |
| trm_phi t = t*} |
|
215 |
||
216 |
ML{*val phi = Morphism.term_morphism trm_phi*}
|
|
217 |
||
218 |
ML{*Morphism.term phi @{term "P x y"}*}
|
|
219 |
||
220 |
text {*
|
|
221 |
@{ML_ind term_morphism in Morphism}
|
|
222 |
||
223 |
@{ML_ind term in Morphism},
|
|
224 |
@{ML_ind thm in Morphism}
|
|
225 |
||
226 |
\begin{readmore}
|
|
227 |
Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
|
|
228 |
\end{readmore}
|
|
229 |
*} |
|
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
230 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
231 |
section {* Misc (TBD) *}
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
232 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
233 |
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
|
234 |
|
|
319
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
235 |
text {*
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
236 |
FIXME: association lists: |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
237 |
@{ML_file "Pure/General/alist.ML"}
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
238 |
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
239 |
FIXME: calling the ML-compiler |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
240 |
|
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
241 |
*} |
|
6bce4acf7f2a
added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents:
318
diff
changeset
|
242 |
|
| 414 | 243 |
section {* What Is In an Isabelle Name? (TBD) *}
|
244 |
||
245 |
text {*
|
|
246 |
On the ML-level of Isabelle, you often have to work with qualified names. |
|
247 |
These are strings with some additional information, such as positional |
|
248 |
information and qualifiers. Such qualified names can be generated with the |
|
249 |
antiquotation @{text "@{binding \<dots>}"}. For example
|
|
250 |
||
251 |
@{ML_response [display,gray]
|
|
252 |
"@{binding \"name\"}"
|
|
253 |
"name"} |
|
254 |
||
255 |
An example where a qualified name is needed is the function |
|
256 |
@{ML_ind define in Local_Theory}. This function is used below to define
|
|
257 |
the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
|
|
258 |
*} |
|
259 |
||
260 |
local_setup %gray {*
|
|
261 |
Local_Theory.define ((@{binding "TrueConj"}, NoSyn),
|
|
262 |
(Attrib.empty_binding, @{term "True \<and> True"})) #> snd *}
|
|
263 |
||
264 |
text {*
|
|
265 |
Now querying the definition you obtain: |
|
266 |
||
267 |
\begin{isabelle}
|
|
268 |
\isacommand{thm}~@{text "TrueConj_def"}\\
|
|
269 |
@{text "> "}~@{thm TrueConj_def}
|
|
270 |
\end{isabelle}
|
|
271 |
||
272 |
\begin{readmore}
|
|
273 |
The basic operations on bindings are implemented in |
|
274 |
@{ML_file "Pure/General/binding.ML"}.
|
|
275 |
\end{readmore}
|
|
276 |
||
277 |
\footnote{\bf FIXME give a better example why bindings are important}
|
|
278 |
\footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
|
|
279 |
why @{ML snd} is needed.}
|
|
280 |
\footnote{\bf FIXME: There should probably a separate section on binding, long-names
|
|
281 |
and sign.} |
|
282 |
||
283 |
*} |
|
284 |
||
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
285 |
|
| 360 | 286 |
ML {* Sign.intern_type @{theory} "list" *}
|
287 |
ML {* Sign.intern_const @{theory} "prod_fun" *}
|
|
288 |
||
| 414 | 289 |
text {*
|
290 |
\footnote{\bf FIXME: Explain the following better; maybe put in a separate
|
|
291 |
section and link with the comment in the antiquotation section.} |
|
292 |
||
293 |
Occasionally you have to calculate what the ``base'' name of a given |
|
| 462 | 294 |
constant is. For this you can use the function @{ML_ind Long_Name.base_name}. For example:
|
| 414 | 295 |
|
| 462 | 296 |
@{ML_response [display,gray] "Long_Name.base_name \"List.list.Nil\"" "\"Nil\""}
|
| 414 | 297 |
|
298 |
\begin{readmore}
|
|
299 |
Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
|
|
300 |
functions about signatures in @{ML_file "Pure/sign.ML"}.
|
|
301 |
\end{readmore}
|
|
302 |
*} |
|
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
303 |
|
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
304 |
text {*
|
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
305 |
@{ML_ind "Binding.name_of"} returns the string without markup
|
| 394 | 306 |
|
307 |
@{ML_ind "Binding.conceal"}
|
|
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
308 |
*} |
|
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
309 |
|
|
388
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
310 |
section {* Concurrency (TBD) *}
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
311 |
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
312 |
text {*
|
|
0b337dedc306
added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents:
387
diff
changeset
|
313 |
@{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
|
314 |
@{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
|
315 |
@{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
|
316 |
*} |
|
387
5dcee4d751ad
completed the unification section
Christian Urban <urbanc@in.tum.de>
parents:
386
diff
changeset
|
317 |
|
| 396 | 318 |
section {* Parse and Print Translations (TBD) *}
|
319 |
||
|
349
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
320 |
section {* Summary *}
|
|
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
321 |
|
|
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
322 |
text {*
|
|
395
2c392f61f400
spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents:
394
diff
changeset
|
323 |
TBD |
|
349
9e374cd891e1
updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents:
348
diff
changeset
|
324 |
*} |
|
318
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
325 |
|
|
efb5fff99c96
split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
326 |
end |