author | Christian Urban <urbanc@in.tum.de> |
Sat, 14 Feb 2009 00:24:05 +0000 | |
changeset 117 | 796c6ea633b3 |
parent 116 | c9ff326e3ce5 |
child 118 | 5f003fdf2653 |
permissions | -rw-r--r-- |
32 | 1 |
theory Ind_Interface |
42
cd612b489504
tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents:
35
diff
changeset
|
2 |
imports "../Base" Simple_Inductive_Package |
32 | 3 |
begin |
4 |
||
113 | 5 |
section {* The Interface \label{sec:ind-interface} *} |
32 | 6 |
|
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
7 |
text {* |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
8 |
The purpose of the package we show next is that the user just specifies the |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
9 |
inductive predicate by stating some introduction rules and then the packages |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
10 |
makes the equivalent definition and derives from it useful properties. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
11 |
To be able to write down the specification in Isabelle, we have to introduce |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
12 |
a new command (see Section~\ref{sec:newcommand}). As the keyword for the new |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
13 |
command we use \simpleinductive{}. The specifications corresponding to our |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
14 |
examples described earlier are: |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
15 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
16 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
17 |
simple_inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
18 |
trcl :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
19 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
20 |
base: "trcl R x x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
21 |
| step: "trcl R x y \<Longrightarrow> R y z \<Longrightarrow> trcl R x z" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
22 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
23 |
simple_inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
24 |
even and odd |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
25 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
26 |
even0: "even 0" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
27 |
| evenS: "odd n \<Longrightarrow> even (Suc n)" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
28 |
| oddS: "even n \<Longrightarrow> odd (Suc n)" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
29 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
30 |
simple_inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
31 |
accpart :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
32 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
33 |
accpartI: "(\<forall>y. R y x \<longrightarrow> accpart R y) \<Longrightarrow> accpart R x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
34 |
|
32 | 35 |
text {* |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
36 |
We expect a constant (or constants) with possible typing annotations and a |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
37 |
list of introduction rules. While these specifications are all |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
38 |
straightforward, there is a technicality we like to deal with to do with |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
39 |
fixed parameters and locales. Remember we pointed out that the parameter |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
40 |
@{text R} is fixed throughout the specifications of @{text trcl} and @{text |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
41 |
accpart}. The point is that they might be fixed in a locale and we like to |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
42 |
support this. Accordingly we treat some parameters of the inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
43 |
definition specially; see Figure~\ref{fig:inddefsfixed} where the transitive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
44 |
closure and accessible part are defined with a fixed parameter @{text R} and |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
45 |
also inside a locale fixing @{text R}. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
46 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
47 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
48 |
text_raw {* |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
49 |
\begin{figure}[p] |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
50 |
\begin{isabelle} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
51 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
52 |
simple_inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
53 |
trcl' for R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
54 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
55 |
base: "trcl' R x x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
56 |
| step: "trcl' R x y \<Longrightarrow> R y z \<Longrightarrow> trcl' R x z" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
57 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
58 |
simple_inductive |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
59 |
accpart' for R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
60 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
61 |
accpartI: "(\<forall>y. R y x \<longrightarrow> accpart' R y) \<Longrightarrow> accpart' R x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
62 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
63 |
locale rel = |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
64 |
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
65 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
66 |
simple_inductive (in rel) trcl'' |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
67 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
68 |
base: "trcl'' x x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
69 |
| step: "trcl'' x y \<Longrightarrow> R y z \<Longrightarrow> trcl'' x z" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
70 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
71 |
simple_inductive (in rel) accpart'' |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
72 |
where |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
73 |
accpartI: "(\<forall>y. R y x \<longrightarrow> accpart'' y) \<Longrightarrow> accpart'' x" |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
74 |
text_raw {* |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
75 |
\end{isabelle} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
76 |
\caption{The first definition is for the transitive closure where the |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
77 |
relation @{text R} is explicitly fixed. Simiraly the second definition |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
78 |
of the accessible part of the relation @{text R}. The last two definitions |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
79 |
specify the same inductive predicates, but this time defined inside |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
80 |
a locale.\label{fig:inddefsfixed}} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
81 |
\end{figure} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
82 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
83 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
84 |
text {* |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
85 |
\begin{figure}[p] |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
86 |
\begin{isabelle} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
87 |
\railnontermfont{\rmfamily\itshape} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
88 |
\railterm{simpleinductive,where,for} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
89 |
\railalias{simpleinductive}{\simpleinductive{}} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
90 |
\railalias{where}{\isacommand{where}} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
91 |
\railalias{for}{\isacommand{for}} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
92 |
\begin{rail} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
93 |
simpleinductive target? fixes (for fixes)? \\ |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
94 |
(where (thmdecl? prop + '|'))? |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
95 |
; |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
96 |
\end{rail} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
97 |
\end{isabelle} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
98 |
\caption{A railroad diagram describing the syntax of \simpleinductive{}. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
99 |
The \emph{target} indicates an optional locale; the \emph{fixes} are an |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
100 |
\isacommand{and}-separated list of names for the inductive predicates (they |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
101 |
can also contain typing- and syntax anotations); similarly the \emph{fixes} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
102 |
after \isacommand{for} to indicate fixed parameters; \emph{prop} stands for a |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
103 |
introduction rule with an optional theorem declaration (\emph{thmdecl}). |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
104 |
\label{fig:railroad}} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
105 |
\end{figure} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
106 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
107 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
108 |
text {* |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
109 |
The syntax of the \simpleinductive{} command can be described by the |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
110 |
railroad diagram in Figure~\ref{fig:railroad}. This diagram more or less |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
111 |
translates directly into the parser |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
112 |
|
117
796c6ea633b3
added an option for linenumbers to the chunk-antiquotation
Christian Urban <urbanc@in.tum.de>
parents:
116
diff
changeset
|
113 |
@{ML_chunk [display,gray,linenos] parser} |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
114 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
115 |
which we also described in Section~\ref{sec:parsingspecs} |
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
116 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
117 |
In order to add a new inductive predicate to a theory with the help of our |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
118 |
package, the user must \emph{invoke} it. For every package, there are |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
119 |
essentially two different ways of invoking it, which we will refer to as |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
120 |
\emph{external} and \emph{internal}. By external invocation we mean that the |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
121 |
package is called from within a theory document. In this case, the |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
122 |
specification of the inductive predicate, including type annotations and |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
123 |
introduction rules, are given as strings by the user. Before the package can |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
124 |
actually make the definition, the type and introduction rules have to be |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
125 |
parsed. In contrast, internal invocation means that the package is called by |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
126 |
some other package. For example, the function definition package |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
127 |
\cite{Krauss-IJCAR06} calls the inductive definition package to define the |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
128 |
graph of the function. However, it is not a good idea for the function |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
129 |
definition package to pass the introduction rules for the function graph to |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
130 |
the inductive definition package as strings. In this case, it is better to |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
131 |
directly pass the rules to the package as a list of terms, which is more |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
132 |
robust than handling strings that are lacking the additional structure of |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
133 |
terms. These two ways of invoking the package are reflected in its ML |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
134 |
programming interface, which consists of two functions: |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
135 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
136 |
|
113 | 137 |
@{ML_chunk [display,gray] SIMPLE_INDUCTIVE_PACKAGE} |
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
138 |
*} |
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
139 |
|
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
140 |
text {* |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
141 |
(FIXME: explain Binding.binding; mixfix; Attrib.src; Attrib.src somewhere else) |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
142 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
143 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
144 |
The function for external invocation of the package is called @{ML |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
145 |
add_inductive in SimpleInductivePackage}, whereas the one for internal |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
146 |
invocation is called @{ML add_inductive_i in SimpleInductivePackage}. Both |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
147 |
of these functions take as arguments the names and types of the inductive |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
148 |
predicates, the names and types of their parameters, the actual introduction |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
149 |
rules and a \emph{local theory}. They return a local theory containing the |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
150 |
definition and the induction principle as well introduction rules. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
151 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
152 |
In contrast to an ordinary theory, which simply consists of a type |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
153 |
signature, as well as tables for constants, axioms and theorems, a local |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
154 |
theory also contains additional context information, such as locally fixed |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
155 |
variables and local assumptions that may be used by the package. The type |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
156 |
@{ML_type local_theory} is identical to the type of \emph{proof contexts} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
157 |
@{ML_type "Proof.context"}, although not every proof context constitutes a |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
158 |
valid local theory. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
159 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
160 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
161 |
Note that @{ML add_inductive_i in SimpleInductivePackage} expects |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
162 |
the types of the predicates and parameters to be specified using the |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
163 |
datatype @{ML_type typ} of Isabelle's logical framework, whereas @{ML |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
164 |
add_inductive in SimpleInductivePackage} expects them to be given as |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
165 |
optional strings. If no string is given for a particular predicate or |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
166 |
parameter, this means that the type should be inferred by the |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
167 |
package. |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
168 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
169 |
|
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
170 |
Additional \emph{mixfix syntax} may be associated with the |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
171 |
predicates and parameters as well. Note that @{ML add_inductive_i in |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
172 |
SimpleInductivePackage} does not allow mixfix syntax to be associated with |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
173 |
parameters, since it can only be used for parsing.\footnote{FIXME: why ist it there then?} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
174 |
The names of the |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
175 |
predicates, parameters and rules are represented by the type @{ML_type |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
176 |
Binding.binding}. Strings can be turned into elements of the type @{ML_type |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
177 |
Binding.binding} using the function @{ML [display] "Binding.name : string -> |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
178 |
Binding.binding"} Each introduction rule is given as a tuple containing its |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
179 |
name, a list of \emph{attributes} and a logical formula. Note that the type |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
180 |
@{ML_type Attrib.binding} used in the list of introduction rules is just a |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
181 |
shorthand for the type @{ML_type "Binding.binding * Attrib.src list"}. The |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
182 |
function @{ML add_inductive_i in SimpleInductivePackage} expects the formula |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
183 |
to be specified using the datatype @{ML_type term}, whereas @{ML |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
184 |
add_inductive in SimpleInductivePackage} expects it to be given as a string. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
185 |
An attribute specifies additional actions and transformations that should be |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
186 |
applied to a theorem, such as storing it in the rule databases used by |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
187 |
automatic tactics like the simplifier. The code of the package, which will |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
188 |
be described in the following section, will mostly treat attributes as a |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
189 |
black box and just forward them to other functions for storing theorems in |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
190 |
local theories. The implementation of the function @{ML add_inductive in |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
191 |
SimpleInductivePackage} for external invocation of the package is quite |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
192 |
simple. Essentially, it just parses the introduction rules and then passes |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
193 |
them on to @{ML add_inductive_i in SimpleInductivePackage}: |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
194 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
195 |
@{ML_chunk [display] add_inductive} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
196 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
197 |
For parsing and type checking the introduction rules, we use the function |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
198 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
199 |
@{ML [display] "Specification.read_specification: |
76
b99fa5fa63fc
adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents:
71
diff
changeset
|
200 |
(Binding.binding * string option * mixfix) list -> (*{variables}*) |
32 | 201 |
(Attrib.binding * string list) list list -> (*{rules}*) |
202 |
local_theory -> |
|
76
b99fa5fa63fc
adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents:
71
diff
changeset
|
203 |
(((Binding.binding * typ) * mixfix) list * |
32 | 204 |
(Attrib.binding * term list) list) * |
205 |
local_theory"} |
|
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
206 |
*} |
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
207 |
|
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
208 |
text {* |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
209 |
During parsing, both predicates and parameters are treated as variables, so |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
210 |
the lists \verb!preds_syn! and \verb!params_syn! are just appended |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
211 |
before being passed to @{ML read_specification in Specification}. Note that the format |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
212 |
for rules supported by @{ML read_specification in Specification} is more general than |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
213 |
what is required for our package. It allows several rules to be associated |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
214 |
with one name, and the list of rules can be partitioned into several |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
215 |
sublists. In order for the list \verb!intro_srcs! of introduction rules |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
216 |
to be acceptable as an input for @{ML read_specification in Specification}, we first |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
217 |
have to turn it into a list of singleton lists. This transformation |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
218 |
has to be reversed later on by applying the function |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
219 |
@{ML [display] "the_single: 'a list -> 'a"} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
220 |
to the list \verb!specs! containing the parsed introduction rules. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
221 |
The function @{ML read_specification in Specification} also returns the list \verb!vars! |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
222 |
of predicates and parameters that contains the inferred types as well. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
223 |
This list has to be chopped into the two lists \verb!preds_syn'! and |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
224 |
\verb!params_syn'! for predicates and parameters, respectively. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
225 |
All variables occurring in a rule but not in the list of variables passed to |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
226 |
@{ML read_specification in Specification} will be bound by a meta-level universal |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
227 |
quantifier. |
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
228 |
*} |
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
229 |
|
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
230 |
text {* |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
231 |
Finally, @{ML read_specification in Specification} also returns another local theory, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
232 |
but we can safely discard it. As an example, let us look at how we can use this |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
233 |
function to parse the introduction rules of the @{text trcl} predicate: |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
234 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
235 |
@{ML_response [display] |
32 | 236 |
"Specification.read_specification |
55 | 237 |
[(Binding.name \"trcl\", NONE, NoSyn), |
238 |
(Binding.name \"r\", SOME \"'a \<Rightarrow> 'a \<Rightarrow> bool\", NoSyn)] |
|
239 |
[[((Binding.name \"base\", []), [\"trcl r x x\"])], |
|
240 |
[((Binding.name \"step\", []), [\"trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z\"])]] |
|
32 | 241 |
@{context}" |
242 |
"((\<dots>, |
|
243 |
[(\<dots>, |
|
244 |
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>), |
|
245 |
Const (\"Trueprop\", \<dots>) $ |
|
246 |
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 0 $ Bound 0))]), |
|
247 |
(\<dots>, |
|
248 |
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>), |
|
249 |
Const (\"all\", \<dots>) $ Abs (\"y\", TFree (\"'a\", \<dots>), |
|
250 |
Const (\"all\", \<dots>) $ Abs (\"z\", TFree (\"'a\", \<dots>), |
|
251 |
Const (\"==>\", \<dots>) $ |
|
252 |
(Const (\"Trueprop\", \<dots>) $ |
|
42
cd612b489504
tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents:
35
diff
changeset
|
253 |
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 2 $ Bound 1)) $ |
32 | 254 |
(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>))))])]), |
255 |
\<dots>) |
|
76
b99fa5fa63fc
adapted to changes in binding.ML
Christian Urban <urbanc@in.tum.de>
parents:
71
diff
changeset
|
256 |
: (((Binding.binding * typ) * mixfix) list * |
32 | 257 |
(Attrib.binding * term list) list) * local_theory"} |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
258 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
259 |
In the list of variables passed to @{ML read_specification in Specification}, we have |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
260 |
used the mixfix annotation @{ML NoSyn} to indicate that we do not want to associate any |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
261 |
mixfix syntax with the variable. Moreover, we have only specified the type of \texttt{r}, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
262 |
whereas the type of \texttt{trcl} is computed using type inference. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
263 |
The local variables \texttt{x}, \texttt{y} and \texttt{z} of the introduction rules |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
264 |
are turned into bound variables with the de Bruijn indices, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
265 |
whereas \texttt{trcl} and \texttt{r} remain free variables. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
266 |
|
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
267 |
*} |
32 | 268 |
|
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
269 |
text {* |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
270 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
271 |
\paragraph{Parsers for theory syntax} |
32 | 272 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
273 |
Although the function @{ML add_inductive in SimpleInductivePackage} parses terms and types, it still |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
274 |
cannot be used to invoke the package directly from within a theory document. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
275 |
In order to do this, we have to write another parser. Before we describe |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
276 |
the process of writing parsers for theory syntax in more detail, we first |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
277 |
show some examples of how we would like to use the inductive definition |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
278 |
package. |
32 | 279 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
280 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
281 |
The definition of the transitive closure should look as follows: |
32 | 282 |
*} |
283 |
||
284 |
text {* |
|
285 |
||
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
286 |
A proposition can be parsed using the function @{ML prop in OuterParse}. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
287 |
Essentially, a proposition is just a string or an identifier, but using the |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
288 |
specific parser function @{ML prop in OuterParse} leads to more instructive |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
289 |
error messages, since the parser will complain that a proposition was expected |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
290 |
when something else than a string or identifier is found. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
291 |
An optional locale target specification of the form \isa{(\isacommand{in}\ $\ldots$)} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
292 |
can be parsed using @{ML opt_target in OuterParse}. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
293 |
The lists of names of the predicates and parameters, together with optional |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
294 |
types and syntax, are parsed using the functions @{ML "fixes" in OuterParse} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
295 |
and @{ML for_fixes in OuterParse}, respectively. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
296 |
In addition, the following function from @{ML_struct SpecParse} for parsing |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
297 |
an optional theorem name and attribute, followed by a delimiter, will be useful: |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
298 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
299 |
\begin{table} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
300 |
@{ML "opt_thm_name: |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
301 |
string -> token list -> Attrib.binding * token list" in SpecParse} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
302 |
\end{table} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
303 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
304 |
We now have all the necessary tools to write the parser for our |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
305 |
\isa{\isacommand{simple{\isacharunderscore}inductive}} command: |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
306 |
|
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
307 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
308 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
309 |
The definition of the parser \verb!ind_decl! closely follows the railroad |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
310 |
diagram shown above. In order to make the code more readable, the structures |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
311 |
@{ML_struct OuterParse} and @{ML_struct OuterKeyword} are abbreviated by |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
312 |
\texttt{P} and \texttt{K}, respectively. Note how the parser combinator |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
313 |
@{ML "!!!" in OuterParse} is used: once the keyword \texttt{where} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
314 |
has been parsed, a non-empty list of introduction rules must follow. |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
315 |
Had we not used the combinator @{ML "!!!" in OuterParse}, a |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
316 |
\texttt{where} not followed by a list of rules would have caused the parser |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
317 |
to respond with the somewhat misleading error message |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
318 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
319 |
\begin{verbatim} |
32 | 320 |
Outer syntax error: end of input expected, but keyword where was found |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
321 |
\end{verbatim} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
322 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
323 |
rather than with the more instructive message |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
324 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
325 |
\begin{verbatim} |
32 | 326 |
Outer syntax error: proposition expected, but terminator was found |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
327 |
\end{verbatim} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
328 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
329 |
Once all arguments of the command have been parsed, we apply the function |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
330 |
@{ML add_inductive in SimpleInductivePackage}, which yields a local theory |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
331 |
transformer of type @{ML_type "local_theory -> local_theory"}. Commands in |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
332 |
Isabelle/Isar are realized by transition transformers of type |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
333 |
@{ML_type [display] "Toplevel.transition -> Toplevel.transition"} |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
334 |
We can turn a local theory transformer into a transition transformer by using |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
335 |
the function |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
336 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
337 |
@{ML [display] "Toplevel.local_theory : string option -> |
32 | 338 |
(local_theory -> local_theory) -> |
339 |
Toplevel.transition -> Toplevel.transition"} |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
340 |
|
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
341 |
which, apart from the local theory transformer, takes an optional name of a locale |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
342 |
to be used as a basis for the local theory. |
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents:
55
diff
changeset
|
343 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
344 |
(FIXME : needs to be adjusted to new parser type) |
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents:
55
diff
changeset
|
345 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
346 |
{\it |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
347 |
The whole parser for our command has type |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
88
diff
changeset
|
348 |
@{text [display] "OuterLex.token list -> |
32 | 349 |
(Toplevel.transition -> Toplevel.transition) * OuterLex.token list"} |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
88
diff
changeset
|
350 |
which is abbreviated by @{text OuterSyntax.parser_fn}. The new command can be added |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
351 |
to the system via the function |
102
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents:
88
diff
changeset
|
352 |
@{text [display] "OuterSyntax.command : |
32 | 353 |
string -> string -> OuterKeyword.T -> OuterSyntax.parser_fn -> unit"} |
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
354 |
which imperatively updates the parser table behind the scenes. } |
60
5b9c6010897b
doem tuning and made the cookbook work again with recent changes (CookBook/Package/Ind_Interface.thy needs to be looked at to see what the problem with the new parser type is)
Christian Urban <urbanc@in.tum.de>
parents:
55
diff
changeset
|
355 |
|
88
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
356 |
In addition to the parser, this |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
357 |
function takes two strings representing the name of the command and a short description, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
358 |
as well as an element of type @{ML_type OuterKeyword.T} describing which \emph{kind} of |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
359 |
command we intend to add. Since we want to add a command for declaring new concepts, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
360 |
we choose the kind @{ML "OuterKeyword.thy_decl"}. Other kinds include |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
361 |
@{ML "OuterKeyword.thy_goal"}, which is similar to @{ML thy_decl in OuterKeyword}, |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
362 |
but requires the user to prove a goal before making the declaration, or |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
363 |
@{ML "OuterKeyword.diag"}, which corresponds to a purely diagnostic command that does |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
364 |
not change the context. For example, the @{ML thy_goal in OuterKeyword} kind is used |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
365 |
by the \isa{\isacommand{function}} command \cite{Krauss-IJCAR06}, which requires the user |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
366 |
to prove that a given set of equations is non-overlapping and covers all cases. The kind |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
367 |
of the command should be chosen with care, since selecting the wrong one can cause strange |
ebbd0dd008c8
adaptation of the package chapter to fit the rest
Christian Urban <urbanc@in.tum.de>
parents:
76
diff
changeset
|
368 |
behaviour of the user interface, such as failure of the undo mechanism. |
32 | 369 |
*} |
370 |
||
371 |
(*<*) |
|
372 |
end |
|
373 |
(*>*) |