author | Christian Urban <urbanc@in.tum.de> |
Tue, 17 Nov 2009 20:36:18 +0100 | |
changeset 391 | ae2f0b40c840 |
parent 365 | 718beb785213 |
child 426 | d94755882e36 |
permissions | -rw-r--r-- |
32 | 1 |
theory Ind_Interface |
346
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents:
316
diff
changeset
|
2 |
imports Ind_Intro 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 |
|
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
7 |
text_raw {* |
296 | 8 |
\begin{figure}[t] |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
9 |
\begin{boxedminipage}{\textwidth} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
10 |
\begin{isabelle} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
11 |
*} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
12 |
simple_inductive |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
13 |
trcl :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
14 |
where |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
15 |
base: "trcl R x x" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
16 |
| step: "trcl R x y \<Longrightarrow> R y z \<Longrightarrow> trcl R x z" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
17 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
18 |
simple_inductive |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
19 |
even and odd |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
20 |
where |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
21 |
even0: "even 0" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
22 |
| evenS: "odd n \<Longrightarrow> even (Suc n)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
23 |
| oddS: "even n \<Longrightarrow> odd (Suc n)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
24 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
25 |
simple_inductive |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
26 |
accpart :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
27 |
where |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
28 |
accpartI: "(\<And>y. R y x \<Longrightarrow> accpart R y) \<Longrightarrow> accpart R x" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
29 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
30 |
(*<*) |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
31 |
datatype trm = |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
32 |
Var "string" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
33 |
| App "trm" "trm" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
34 |
| Lam "string" "trm" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
35 |
(*>*) |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
36 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
37 |
simple_inductive |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
38 |
fresh :: "string \<Rightarrow> trm \<Rightarrow> bool" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
39 |
where |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
40 |
fresh_var: "a\<noteq>b \<Longrightarrow> fresh a (Var b)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
41 |
| fresh_app: "\<lbrakk>fresh a t; fresh a s\<rbrakk> \<Longrightarrow> fresh a (App t s)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
42 |
| fresh_lam1: "fresh a (Lam a t)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
43 |
| fresh_lam2: "\<lbrakk>a\<noteq>b; fresh a t\<rbrakk> \<Longrightarrow> fresh a (Lam b t)" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
44 |
text_raw {* |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
45 |
\end{isabelle} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
46 |
\end{boxedminipage} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
47 |
\caption{Specification given by the user for the inductive predicates |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
48 |
@{term "trcl"}, @{term "even"} and @{term "odd"}, @{term "accpart"} and |
244
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
49 |
@{term "fresh"}.\label{fig:specs}} |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
50 |
\end{figure} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
51 |
*} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
52 |
|
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
53 |
text {* |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
54 |
To be able to write down the specifications of inductive predicates, we have |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
55 |
to introduce a new command (see Section~\ref{sec:newcommand}). As the |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
56 |
keyword for the new command we chose \simpleinductive{}. Examples of |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
57 |
specifications from the previous section are shown in |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
58 |
Figure~\ref{fig:specs}. The syntax used in these examples more or |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
59 |
less translates directly into the parser: |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
60 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
61 |
*} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
62 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
63 |
ML{*val spec_parser = |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
64 |
OuterParse.fixes -- |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
65 |
Scan.optional |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
66 |
(OuterParse.$$$ "where" |-- |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
67 |
OuterParse.!!! |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
68 |
(OuterParse.enum1 "|" |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
69 |
(SpecParse.opt_thm_name ":" -- OuterParse.prop))) []*} |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
70 |
|
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
71 |
text {* |
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
72 |
which we explained in Section~\ref{sec:parsingspecs}. There is no code included |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
73 |
for parsing the keyword and what is called a \emph{target}. The latter can be given |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
74 |
optionally after the keyword. The target is an ``advanced'' feature which we will |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
75 |
inherit for ``free'' from the infrastructure on which we shall build the package. |
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
76 |
The target stands for a locale and allows us to specify |
127
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
77 |
*} |
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
78 |
|
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
79 |
locale rel = |
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
80 |
fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" |
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
81 |
|
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
82 |
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
|
83 |
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
|
84 |
locale as follows: |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
85 |
*} |
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
86 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
127
diff
changeset
|
87 |
simple_inductive (in rel) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
127
diff
changeset
|
88 |
trcl' |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
89 |
where |
127
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
90 |
base: "trcl' x x" |
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
91 |
| 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
|
92 |
|
129
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
127
diff
changeset
|
93 |
simple_inductive (in rel) |
e0d368a45537
started a section about simprocs
Christian Urban <urbanc@in.tum.de>
parents:
127
diff
changeset
|
94 |
accpart' |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
95 |
where |
127
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
96 |
accpartI: "(\<And>y. R y x \<Longrightarrow> accpart' y) \<Longrightarrow> accpart' x" |
244
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
97 |
(*<*)ML %no{*fun filtered_input str = |
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
98 |
filter OuterLex.is_proper (OuterSyntax.scan Position.none str) |
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
99 |
fun parse p input = Scan.finite OuterLex.stopper (Scan.error p) input*}(*>*) |
116
c9ff326e3ce5
more changes to the package chapter
Christian Urban <urbanc@in.tum.de>
parents:
113
diff
changeset
|
100 |
text {* |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
101 |
Note that in these definitions the parameter @{text R}, standing for the |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
102 |
relation, is left implicit. For the moment we will ignore this kind |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
103 |
of implicit parameters and rely on the fact that the infrastructure will |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
104 |
deal with them. Later, however, we will come back to them. |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
105 |
|
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
106 |
If we feed into the parser the string that corresponds to our definition |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
107 |
of @{term even} and @{term odd} |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
108 |
|
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
109 |
@{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
|
110 |
"let |
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
111 |
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
|
112 |
(\"even and odd \" ^ |
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
113 |
\"where \" ^ |
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
114 |
\" 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
|
115 |
\"| evenS[intro]: \\\"odd n \<Longrightarrow> even (Suc n)\\\" \" ^ |
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
116 |
\"| oddS[intro]: \\\"even n \<Longrightarrow> odd (Suc n)\\\" \") |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
117 |
in |
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
118 |
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
|
119 |
end" |
186
371e4375c994
made the Ackermann function example safer and included suggestions from MW
Christian Urban <urbanc@in.tum.de>
parents:
183
diff
changeset
|
120 |
"(([(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
|
121 |
[((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
|
122 |
((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
|
123 |
((oddS,\<dots>), \"\\^E\\^Ftoken\\^Eeven n \<Longrightarrow> odd (Suc n)\\^E\\^F\\^E\")]), [])"} |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
124 |
|
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
125 |
then we get back the specifications of the predicates (with type and syntax annotations), |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
126 |
and specifications of the introduction rules. This is all the information we |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
127 |
need for calling the package and setting up the keyword. The latter is |
295
24c68350d059
polished the package chapter used FOCUS to explain the subproofs
Christian Urban <urbanc@in.tum.de>
parents:
256
diff
changeset
|
128 |
done in Lines 5 to 7 in the code below. |
120
c39f83d8daeb
some polishing; split up the file External Solver into two
Christian Urban <urbanc@in.tum.de>
parents:
118
diff
changeset
|
129 |
*} |
245
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
parents:
244
diff
changeset
|
130 |
(*<*)ML %no{*fun add_inductive_cmd pred_specs rule_specs lthy = lthy |
53112deda119
Jasmin and Christian added examples for the pretty-printing section
Christian Urban <urbanc@in.tum.de>
parents:
244
diff
changeset
|
131 |
fun add_inductive pred_specs rule_specs lthy = lthy*}(*>*) |
365 | 132 |
ML_val %linenosgray{*val specification : (local_theory -> local_theory) parser = |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
133 |
spec_parser >> |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
134 |
(fn (pred_specs, rule_specs) => add_inductive_cmd pred_specs rule_specs) |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
135 |
|
244
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
136 |
val _ = OuterSyntax.local_theory "simple_inductive2" |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
137 |
"definition of simple inductive predicates" |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
138 |
OuterKeyword.thy_decl specification*} |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
139 |
|
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
140 |
text {* |
316
74f0a06f751f
further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents:
315
diff
changeset
|
141 |
We call @{ML_ind local_theory in OuterSyntax} with the kind-indicator |
74f0a06f751f
further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents:
315
diff
changeset
|
142 |
@{ML_ind thy_decl in OuterKeyword} since the package does not need to open |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
143 |
up any proof (see Section~\ref{sec:newcommand}). |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
144 |
The auxiliary function @{text specification} in Lines 1 to 3 |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
145 |
gathers the information from the parser to be processed further |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
146 |
by the function @{text "add_inductive_cmd"}, which we describe below. |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
147 |
|
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
148 |
Note that the predicates when they come out of the parser are just some |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
149 |
``naked'' strings: they have no type yet (even if we annotate them with |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
150 |
types) and they are also not defined constants yet (which the predicates |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
151 |
eventually will be). Also the introduction rules are just strings. What we have |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
152 |
to do first is to transform the parser's output into some internal |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
153 |
datastructures that can be processed further. For this we can use the |
316
74f0a06f751f
further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents:
315
diff
changeset
|
154 |
function @{ML_ind read_spec in Specification}. This function takes some strings |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
155 |
(with possible typing annotations) and some rule specifications, and attempts |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
156 |
to find a typing according to the given type constraints given by the |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
157 |
user and the type constraints by the ``ambient'' theory. It returns |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
158 |
the type for the predicates and also returns typed terms for the |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
159 |
introduction rules. So at the heart of the function |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
160 |
@{text "add_inductive_cmd"} is a call to @{ML read_spec in Specification}. |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
161 |
*} |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
162 |
|
244
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
163 |
ML_val{*fun add_inductive_cmd pred_specs rule_specs lthy = |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
164 |
let |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
165 |
val ((pred_specs', rule_specs'), _) = |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
166 |
Specification.read_spec pred_specs rule_specs lthy |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
167 |
in |
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
168 |
add_inductive pred_specs' rule_specs' lthy |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
169 |
end*} |
53
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
170 |
|
0c3580c831a4
removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents:
42
diff
changeset
|
171 |
text {* |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
172 |
Once we have the input data as some internal datastructure, we call |
244
dc95a56b1953
fixed the problem with double definition of even and odd
Christian Urban <urbanc@in.tum.de>
parents:
224
diff
changeset
|
173 |
the function @{text add_inductive}. This function does the heavy duty |
224
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
174 |
lifting in the package: it generates definitions for the |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
175 |
predicates and derives from them corresponding induction principles and |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
176 |
introduction rules. The description of this function will span over |
647cab4a72c2
finished the heavy duty stuff for the inductive package
Christian Urban <urbanc@in.tum.de>
parents:
219
diff
changeset
|
177 |
the next two sections. |
127
74846cb0fff9
updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents:
126
diff
changeset
|
178 |
*} |
219
98d43270024f
more work on the simple inductive chapter
Christian Urban <urbanc@in.tum.de>
parents:
218
diff
changeset
|
179 |
(*<*)end(*>*) |