| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 31 Mar 2009 20:31:18 +0100 | |
| changeset 218 | 7ff7325e3b4e | 
| parent 215 | 8d1a344a621e | 
| child 219 | 98d43270024f | 
| permissions | -rw-r--r-- | 
| 32 | 1  | 
theory Ind_Interface  | 
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
2  | 
imports "../Base" "../Parsing" Ind_Prelims Simple_Inductive_Package  | 
| 32 | 3  | 
begin  | 
4  | 
||
| 
215
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
5  | 
section {* Parsing and Typing the Specification\label{sec:interface} *}
 | 
| 
124
 
0b9fa606a746
added to the first-steps section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
120 
diff
changeset
 | 
6  | 
|
| 
 
0b9fa606a746
added to the first-steps section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
120 
diff
changeset
 | 
7  | 
text {* 
 | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
8  | 
To be able to write down the specification in Isabelle, we have to introduce  | 
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
9  | 
  a new command (see Section~\ref{sec:newcommand}).  As the keyword for the
 | 
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
10  | 
  new command we chose \simpleinductive{}. The ``infrastructure'' already 
 | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
11  | 
provides an ``advanced'' feature for this command. For example we will  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
12  | 
be able to declare the locale  | 
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
13  | 
*}  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
14  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
15  | 
locale rel =  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
16  | 
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
17  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
18  | 
text {*
 | 
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
19  | 
and then define the transitive closure and the accessible part of this  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
20  | 
locale as follows:  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
21  | 
*}  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
22  | 
|
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
23  | 
simple_inductive (in rel)  | 
| 
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
24  | 
trcl'  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
25  | 
where  | 
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
26  | 
base: "trcl' x x"  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
27  | 
| step: "trcl' x y \<Longrightarrow> R y z \<Longrightarrow> trcl' x z"  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
28  | 
|
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
29  | 
simple_inductive (in rel)  | 
| 
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
30  | 
accpart'  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
31  | 
where  | 
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
32  | 
accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x"  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
33  | 
|
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
34  | 
text {*
 | 
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
35  | 
  \begin{figure}[t]
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
36  | 
  \begin{isabelle}
 | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
37  | 
  \railnontermfont{\rmfamily\itshape}
 | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
38  | 
  \railterm{simpleinductive,where,for}
 | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
39  | 
  \railalias{simpleinductive}{\simpleinductive{}}
 | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
40  | 
  \railalias{where}{\isacommand{where}}
 | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
41  | 
  \railalias{for}{\isacommand{for}}
 | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
42  | 
  \begin{rail}
 | 
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
43  | 
simpleinductive target?\\ fixes  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
44  | 
(where (thmdecl? prop + '|'))?  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
45  | 
;  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
46  | 
  \end{rail}
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
47  | 
  \end{isabelle}
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
48  | 
  \caption{A railroad diagram describing the syntax of \simpleinductive{}. 
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
49  | 
  The \emph{target} indicates an optional locale; the \emph{fixes} are an 
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
50  | 
  \isacommand{and}-separated list of names for the inductive predicates (they
 | 
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
51  | 
  can also contain typing- and syntax anotations); \emph{prop} stands for a 
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
52  | 
  introduction rule with an optional theorem declaration (\emph{thmdecl}).
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
53  | 
  \label{fig:railroad}}
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
54  | 
  \end{figure}
 | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
55  | 
*}  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
56  | 
|
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
57  | 
text {*
 | 
| 
129
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
58  | 
This leads directly to the railroad diagram shown in  | 
| 
 
e0d368a45537
started a section about simprocs
 
Christian Urban <urbanc@in.tum.de> 
parents: 
127 
diff
changeset
 | 
59  | 
  Figure~\ref{fig:railroad} for the syntax of \simpleinductive{}. This diagram
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
60  | 
more or less translates directly into the parser:  | 
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
61  | 
*}  | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
62  | 
|
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
63  | 
ML %linenosgray{*val spec_parser = 
 | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
64  | 
OuterParse.fixes --  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
65  | 
Scan.optional  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
66  | 
(OuterParse.$$$ "where" |--  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
67  | 
OuterParse.!!!  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
68  | 
(OuterParse.enum1 "|"  | 
| 
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
69  | 
(SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*}  | 
| 
118
 
5f003fdf2653
polished and added more material to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
117 
diff
changeset
 | 
70  | 
|
| 
218
 
7ff7325e3b4e
started to adapt the rest of chapter 5 to the simplified version without parameters (they will be described in the extension section)
 
Christian Urban <urbanc@in.tum.de> 
parents: 
215 
diff
changeset
 | 
71  | 
text {*
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
72  | 
  which we described in Section~\ref{sec:parsingspecs}. If we feed into the 
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
73  | 
  parser the string (which corresponds to our definition of @{term even} and 
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
74  | 
  @{term odd}):
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
75  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
76  | 
  @{ML_response [display,gray]
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
77  | 
"let  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
78  | 
val input = filtered_input  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
79  | 
(\"even and odd \" ^  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
80  | 
\"where \" ^  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
81  | 
\" even0[intro]: \\\"even 0\\\" \" ^  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
82  | 
\"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
83  | 
\"| oddS[intro]: \\\"even n \<Longrightarrow> odd (Suc n)\\\"\")  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
84  | 
in  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
85  | 
parse spec_parser input  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
86  | 
end"  | 
| 
186
 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 
Christian Urban <urbanc@in.tum.de> 
parents: 
183 
diff
changeset
 | 
87  | 
"(([(even, NONE, NoSyn), (odd, NONE, NoSyn)],  | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
88  | 
[((even0,\<dots>), \"\\^E\\^Ftoken\\^Eeven 0\\^E\\^F\\^E\"),  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
89  | 
((evenS,\<dots>), \"\\^E\\^Ftoken\\^Eodd n \<Longrightarrow> even (Suc n)\\^E\\^F\\^E\"),  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
90  | 
((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"}  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
91  | 
*}  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
92  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
93  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
94  | 
text {*
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
95  | 
  then we get back a locale (in this case @{ML NONE}), the predicates (with type
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
96  | 
and syntax annotations), the parameters (similar as the predicates) and  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
97  | 
the specifications of the introduction rules.  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
98  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
99  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
100  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
101  | 
This is all the information we  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
102  | 
need for calling the package and setting up the keyword. The latter is  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
103  | 
done in Lines 6 and 7 in the code below.  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
104  | 
|
| 
118
 
5f003fdf2653
polished and added more material to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
117 
diff
changeset
 | 
105  | 
  @{ML_chunk [display,gray,linenos] syntax}
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
106  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
107  | 
  We call @{ML OuterSyntax.command} with the kind-indicator @{ML
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
108  | 
OuterKeyword.thy_decl} since the package does not need to open up any goal  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
109  | 
  state (see Section~\ref{sec:newcommand}). Note that the predicates and
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
110  | 
parameters are at the moment only some ``naked'' variables: they have no  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
111  | 
type yet (even if we annotate them with types) and they are also no defined  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
112  | 
constants yet (which the predicates will eventually be). In Lines 1 to 4 we  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
113  | 
gather the information from the parser to be processed further. The locale  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
114  | 
  is passed as argument to the function @{ML
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
115  | 
  Toplevel.local_theory}.\footnote{FIXME Is this already described?} The other
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
116  | 
arguments, i.e.~the predicates, parameters and intro rule specifications,  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
117  | 
  are passed to the function @{ML add_inductive in SimpleInductivePackage}
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
118  | 
(Line 4).  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
119  | 
|
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
120  | 
We now come to the second subtask of the package, namely transforming the  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
121  | 
parser output into some internal datastructures that can be processed further.  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
122  | 
Remember that at the moment the introduction rules are just strings, and even  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
123  | 
if the predicates and parameters can contain some typing annotations, they  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
124  | 
are not yet in any way reflected in the introduction rules. So the task of  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
125  | 
  @{ML add_inductive in SimpleInductivePackage} is to transform the strings
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
126  | 
into properly typed terms. For this it can use the function  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
127  | 
  @{ML read_spec in Specification}. This function takes some constants
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
128  | 
with possible typing annotations and some rule specifications and attempts to  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
129  | 
find a type according to the given type constraints and the type constraints  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
130  | 
by the surrounding (local theory). However this function is a bit  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
131  | 
too general for our purposes: we want that each introduction rule has only  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
132  | 
  name (for example @{text even0} or @{text evenS}), if a name is given at all.
 | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
133  | 
  The function @{ML read_spec in Specification} however allows more
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
134  | 
than one rule. Since it is quite convenient to rely on this function (instead of  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
135  | 
building your own) we just quick ly write a wrapper function that translates  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
136  | 
between our specific format and the general format expected by  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
137  | 
  @{ML read_spec in Specification}. The code of this wrapper is as follows:
 | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
138  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
139  | 
  @{ML_chunk [display,gray,linenos] read_specification}
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
140  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
141  | 
It takes a list of constants, a list of rule specifications and a local theory  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
142  | 
as input. Does the transformation of the rule specifications in Line 3; calls  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
143  | 
the function and transforms the now typed rule specifications back into our  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
144  | 
format and returns the type parameter and typed rule specifications.  | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
145  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
146  | 
|
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
147  | 
   @{ML_chunk [display,gray,linenos] add_inductive}
 | 
| 
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
148  | 
|
| 
53
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
149  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
150  | 
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
 | 
151  | 
  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
 | 
152  | 
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
 | 
153  | 
  \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
 | 
154  | 
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
 | 
155  | 
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
 | 
156  | 
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
 | 
157  | 
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
 | 
158  | 
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
 | 
159  | 
some other package. For example, the function definition package  | 
| 
120
 
c39f83d8daeb
some polishing; split up the file External Solver into two
 
Christian Urban <urbanc@in.tum.de> 
parents: 
118 
diff
changeset
 | 
160  | 
calls the inductive definition package to define the  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
161  | 
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
 | 
162  | 
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
 | 
163  | 
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
 | 
164  | 
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
 | 
165  | 
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
 | 
166  | 
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
 | 
167  | 
programming interface, which consists of two functions:  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
168  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
169  | 
|
| 113 | 170  | 
  @{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
 | 
171  | 
*}  | 
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
172  | 
|
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
173  | 
text {*
 | 
| 
124
 
0b9fa606a746
added to the first-steps section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
120 
diff
changeset
 | 
174  | 
(FIXME: explain Binding.binding; Attrib.binding somewhere else)  | 
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
175  | 
|
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
176  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
177  | 
  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
 | 
178  | 
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
 | 
179  | 
  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
 | 
180  | 
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
 | 
181  | 
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
 | 
182  | 
  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
 | 
183  | 
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
 | 
184  | 
|
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
185  | 
  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
 | 
186  | 
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
 | 
187  | 
  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
 | 
188  | 
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
 | 
189  | 
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
 | 
190  | 
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
 | 
191  | 
package.  | 
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
192  | 
|
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
193  | 
|
| 
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
194  | 
  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
 | 
195  | 
  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
 | 
196  | 
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
 | 
197  | 
  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
 | 
198  | 
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
 | 
199  | 
  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
 | 
200  | 
  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
 | 
201  | 
  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
 | 
202  | 
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
 | 
203  | 
  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
 | 
204  | 
  @{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
 | 
205  | 
  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
 | 
206  | 
  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
 | 
207  | 
  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
 | 
208  | 
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
 | 
209  | 
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
 | 
210  | 
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
 | 
211  | 
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
 | 
212  | 
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
 | 
213  | 
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
 | 
214  | 
  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
 | 
215  | 
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
 | 
216  | 
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
 | 
217  | 
  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
 | 
218  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
219  | 
  @{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
 | 
220  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
221  | 
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
 | 
222  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
223  | 
  @{ML [display] "Specification.read_specification:
 | 
| 
76
 
b99fa5fa63fc
adapted to changes in binding.ML
 
Christian Urban <urbanc@in.tum.de> 
parents: 
71 
diff
changeset
 | 
224  | 
  (Binding.binding * string option * mixfix) list ->  (*{variables}*)
 | 
| 
176
 
3da5f3f07d8b
updated to new read_specification
 
Christian Urban <urbanc@in.tum.de> 
parents: 
129 
diff
changeset
 | 
225  | 
  (Attrib.binding * string list) list ->  (*{rules}*)
 | 
| 32 | 226  | 
local_theory ->  | 
| 
76
 
b99fa5fa63fc
adapted to changes in binding.ML
 
Christian Urban <urbanc@in.tum.de> 
parents: 
71 
diff
changeset
 | 
227  | 
(((Binding.binding * typ) * mixfix) list *  | 
| 32 | 228  | 
(Attrib.binding * term list) list) *  | 
229  | 
local_theory"}  | 
|
| 
53
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
230  | 
*}  | 
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
231  | 
|
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
232  | 
text {*
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
233  | 
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
 | 
234  | 
the lists \verb!preds_syn! and \verb!params_syn! are just appended  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
235  | 
  before being passed to @{ML read_spec in Specification}. Note that the format
 | 
| 
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
236  | 
  for rules supported by @{ML read_spec in Specification} is more general than
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
237  | 
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
 | 
238  | 
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
 | 
239  | 
sublists. In order for the list \verb!intro_srcs! of introduction rules  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
240  | 
  to be acceptable as an input for @{ML read_spec in Specification}, we first
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
241  | 
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
 | 
242  | 
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
 | 
243  | 
  @{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
 | 
244  | 
to the list \verb!specs! containing the parsed introduction rules.  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
245  | 
  The function @{ML read_spec in Specification} also returns the list \verb!vars!
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
246  | 
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
 | 
247  | 
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
 | 
248  | 
\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
 | 
249  | 
All variables occurring in a rule but not in the list of variables passed to  | 
| 
183
 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 
Christian Urban <urbanc@in.tum.de> 
parents: 
178 
diff
changeset
 | 
250  | 
  @{ML read_spec in Specification} will be bound by a meta-level universal
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
251  | 
quantifier.  | 
| 
53
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
252  | 
*}  | 
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
253  | 
|
| 
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
254  | 
text {*
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
255  | 
  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
 | 
256  | 
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
 | 
257  | 
  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
 | 
258  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
259  | 
  @{ML_response [display]
 | 
| 32 | 260  | 
"Specification.read_specification  | 
| 55 | 261  | 
[(Binding.name \"trcl\", NONE, NoSyn),  | 
262  | 
(Binding.name \"r\", SOME \"'a \<Rightarrow> 'a \<Rightarrow> bool\", NoSyn)]  | 
|
| 
176
 
3da5f3f07d8b
updated to new read_specification
 
Christian Urban <urbanc@in.tum.de> 
parents: 
129 
diff
changeset
 | 
263  | 
[((Binding.name \"base\", []), [\"trcl r x x\"]),  | 
| 
 
3da5f3f07d8b
updated to new read_specification
 
Christian Urban <urbanc@in.tum.de> 
parents: 
129 
diff
changeset
 | 
264  | 
((Binding.name \"step\", []), [\"trcl r x y \<Longrightarrow> r y z \<Longrightarrow> trcl r x z\"])]  | 
| 32 | 265  | 
  @{context}"
 | 
266  | 
"((\<dots>,  | 
|
267  | 
[(\<dots>,  | 
|
268  | 
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),  | 
|
269  | 
Const (\"Trueprop\", \<dots>) $  | 
|
270  | 
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 0 $ Bound 0))]),  | 
|
271  | 
(\<dots>,  | 
|
272  | 
[Const (\"all\", \<dots>) $ Abs (\"x\", TFree (\"'a\", \<dots>),  | 
|
273  | 
Const (\"all\", \<dots>) $ Abs (\"y\", TFree (\"'a\", \<dots>),  | 
|
274  | 
Const (\"all\", \<dots>) $ Abs (\"z\", TFree (\"'a\", \<dots>),  | 
|
275  | 
Const (\"==>\", \<dots>) $  | 
|
276  | 
(Const (\"Trueprop\", \<dots>) $  | 
|
| 
42
 
cd612b489504
tuned mostly antiquotation and text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
35 
diff
changeset
 | 
277  | 
(Free (\"trcl\", \<dots>) $ Free (\"r\", \<dots>) $ Bound 2 $ Bound 1)) $  | 
| 32 | 278  | 
(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>))))])]),  | 
279  | 
\<dots>)  | 
|
| 
76
 
b99fa5fa63fc
adapted to changes in binding.ML
 
Christian Urban <urbanc@in.tum.de> 
parents: 
71 
diff
changeset
 | 
280  | 
: (((Binding.binding * typ) * mixfix) list *  | 
| 32 | 281  | 
(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
 | 
282  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
283  | 
  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
 | 
284  | 
  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
 | 
285  | 
  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
 | 
286  | 
  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
 | 
287  | 
  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
 | 
288  | 
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
 | 
289  | 
  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
 | 
290  | 
|
| 
53
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
291  | 
*}  | 
| 32 | 292  | 
|
| 
53
 
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
 
Christian Urban <urbanc@in.tum.de> 
parents: 
42 
diff
changeset
 | 
293  | 
text {*
 | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
294  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
295  | 
  \paragraph{Parsers for theory syntax}
 | 
| 32 | 296  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
297  | 
  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
 | 
298  | 
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
 | 
299  | 
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
 | 
300  | 
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
 | 
301  | 
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
 | 
302  | 
package.  | 
| 32 | 303  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
304  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
305  | 
The definition of the transitive closure should look as follows:  | 
| 32 | 306  | 
*}  | 
307  | 
||
| 
178
 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
177 
diff
changeset
 | 
308  | 
ML {* SpecParse.opt_thm_name *}
 | 
| 
 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
177 
diff
changeset
 | 
309  | 
|
| 32 | 310  | 
text {*
 | 
311  | 
||
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
312  | 
  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
 | 
313  | 
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
 | 
314  | 
  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
 | 
315  | 
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
 | 
316  | 
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
 | 
317  | 
  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
 | 
318  | 
  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
 | 
319  | 
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
 | 
320  | 
  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
 | 
321  | 
  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
 | 
322  | 
  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
 | 
323  | 
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
 | 
324  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
325  | 
  \begin{table}
 | 
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
326  | 
  @{ML "opt_thm_name:
 | 
| 
178
 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
177 
diff
changeset
 | 
327  | 
string -> Attrib.binding parser" in SpecParse}  | 
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
328  | 
  \end{table}
 | 
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
329  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
330  | 
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
 | 
331  | 
  \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
 | 
332  | 
|
| 
116
 
c9ff326e3ce5
more changes to the package chapter
 
Christian Urban <urbanc@in.tum.de> 
parents: 
113 
diff
changeset
 | 
333  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
334  | 
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
 | 
335  | 
  @{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
 | 
336  | 
  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
 | 
337  | 
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
 | 
338  | 
  @{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
 | 
339  | 
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
 | 
340  | 
the function  | 
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
341  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
342  | 
  @{ML [display] "Toplevel.local_theory : string option ->
 | 
| 32 | 343  | 
(local_theory -> local_theory) ->  | 
344  | 
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
 | 
345  | 
|
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
346  | 
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
 | 
347  | 
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
 | 
348  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
349  | 
(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
 | 
350  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
351  | 
  {\it
 | 
| 
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
352  | 
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
 | 
353  | 
  @{text [display] "OuterLex.token list ->
 | 
| 32 | 354  | 
(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
 | 
355  | 
  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
 | 
356  | 
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
 | 
357  | 
  @{text [display] "OuterSyntax.command :
 | 
| 32 | 358  | 
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
 | 
359  | 
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
 | 
360  | 
|
| 
88
 
ebbd0dd008c8
adaptation of the package chapter to fit the rest
 
Christian Urban <urbanc@in.tum.de> 
parents: 
76 
diff
changeset
 | 
361  | 
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
 | 
362  | 
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
 | 
363  | 
  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
 | 
364  | 
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
 | 
365  | 
  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
 | 
366  | 
  @{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
 | 
367  | 
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
 | 
368  | 
  @{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
 | 
369  | 
  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
 | 
370  | 
  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
 | 
371  | 
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
 | 
372  | 
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
 | 
373  | 
behaviour of the user interface, such as failure of the undo mechanism.  | 
| 32 | 374  | 
*}  | 
375  | 
||
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
376  | 
text {*
 | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
377  | 
  Note that the @{text trcl} predicate has two different kinds of parameters: the
 | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
378  | 
  first parameter @{text R} stays \emph{fixed} throughout the definition, whereas
 | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
379  | 
the second and third parameter changes in the ``recursive call''. This will  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
380  | 
become important later on when we deal with fixed parameters and locales.  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
381  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
382  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
383  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
384  | 
The purpose of the package we show next is that the user just specifies the  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
385  | 
inductive predicate by stating some introduction rules and then the packages  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
386  | 
makes the equivalent definition and derives from it the needed properties.  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
387  | 
*}  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
388  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
389  | 
text {*
 | 
| 212 | 390  | 
(FIXME: perhaps move somewhere else)  | 
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
391  | 
|
| 212 | 392  | 
The point of these examples is to get a feeling what the automatic proofs  | 
393  | 
should do in order to solve all inductive definitions we throw at them. For this  | 
|
394  | 
it is instructive to look at the general construction principle  | 
|
395  | 
of inductive definitions, which we shall do in the next section.  | 
|
| 
215
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
396  | 
|
| 
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
397  | 
The code of the inductive package falls roughly in tree parts: the first  | 
| 
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
398  | 
deals with the definitions, the second with the induction principles and  | 
| 
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
399  | 
the third with the introduction rules.  | 
| 
 
8d1a344a621e
more work on the inductive package
 
Christian Urban <urbanc@in.tum.de> 
parents: 
212 
diff
changeset
 | 
400  | 
|
| 
127
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
401  | 
*}  | 
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
402  | 
|
| 
 
74846cb0fff9
updated and added two tentative recipes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
126 
diff
changeset
 | 
403  | 
|
| 32 | 404  | 
(*<*)  | 
405  | 
end  | 
|
406  | 
(*>*)  |