author | boehmes |
Sat, 03 Oct 2009 12:10:16 +0200 | |
changeset 32865 | f8d1e16ec758 |
parent 30168 | 9a20be5be90b |
child 40255 | 9ffbc25e1606 |
permissions | -rw-r--r-- |
26840 | 1 |
theory ZF_Specific |
26894 | 2 |
imports Main |
26840 | 3 |
begin |
4 |
||
26852 | 5 |
chapter {* Isabelle/ZF \label{ch:zf} *} |
26845 | 6 |
|
7 |
section {* Type checking *} |
|
8 |
||
9 |
text {* |
|
10 |
The ZF logic is essentially untyped, so the concept of ``type |
|
11 |
checking'' is performed as logical reasoning about set-membership |
|
12 |
statements. A special method assists users in this task; a version |
|
13 |
of this is already declared as a ``solver'' in the standard |
|
14 |
Simplifier setup. |
|
15 |
||
16 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
17 |
@{command_def (ZF) "print_tcset"}@{text "\<^sup>*"} & : & @{text "context \<rightarrow>"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
18 |
@{method_def (ZF) typecheck} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
19 |
@{attribute_def (ZF) TC} & : & @{text attribute} \\ |
26845 | 20 |
\end{matharray} |
21 |
||
22 |
\begin{rail} |
|
23 |
'TC' (() | 'add' | 'del') |
|
24 |
; |
|
25 |
\end{rail} |
|
26 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
27 |
\begin{description} |
26845 | 28 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
29 |
\item @{command (ZF) "print_tcset"} prints the collection of |
26845 | 30 |
typechecking rules of the current context. |
31 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
32 |
\item @{method (ZF) typecheck} attempts to solve any pending |
26845 | 33 |
type-checking problems in subgoals. |
34 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
35 |
\item @{attribute (ZF) TC} adds or deletes type-checking rules from |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
36 |
the context. |
26845 | 37 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
38 |
\end{description} |
26845 | 39 |
*} |
40 |
||
41 |
||
42 |
section {* (Co)Inductive sets and datatypes *} |
|
43 |
||
44 |
subsection {* Set definitions *} |
|
45 |
||
46 |
text {* |
|
47 |
In ZF everything is a set. The generic inductive package also |
|
48 |
provides a specific view for ``datatype'' specifications. |
|
49 |
Coinductive definitions are available in both cases, too. |
|
50 |
||
51 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
52 |
@{command_def (ZF) "inductive"} & : & @{text "theory \<rightarrow> theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
53 |
@{command_def (ZF) "coinductive"} & : & @{text "theory \<rightarrow> theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
54 |
@{command_def (ZF) "datatype"} & : & @{text "theory \<rightarrow> theory"} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
55 |
@{command_def (ZF) "codatatype"} & : & @{text "theory \<rightarrow> theory"} \\ |
26845 | 56 |
\end{matharray} |
57 |
||
58 |
\begin{rail} |
|
59 |
('inductive' | 'coinductive') domains intros hints |
|
60 |
; |
|
61 |
||
62 |
domains: 'domains' (term + '+') ('<=' | subseteq) term |
|
63 |
; |
|
64 |
intros: 'intros' (thmdecl? prop +) |
|
65 |
; |
|
66 |
hints: monos? condefs? typeintros? typeelims? |
|
67 |
; |
|
68 |
monos: ('monos' thmrefs)? |
|
69 |
; |
|
70 |
condefs: ('con\_defs' thmrefs)? |
|
71 |
; |
|
72 |
typeintros: ('type\_intros' thmrefs)? |
|
73 |
; |
|
74 |
typeelims: ('type\_elims' thmrefs)? |
|
75 |
; |
|
76 |
\end{rail} |
|
77 |
||
78 |
In the following syntax specification @{text "monos"}, @{text |
|
79 |
typeintros}, and @{text typeelims} are the same as above. |
|
80 |
||
81 |
\begin{rail} |
|
82 |
('datatype' | 'codatatype') domain? (dtspec + 'and') hints |
|
83 |
; |
|
84 |
||
85 |
domain: ('<=' | subseteq) term |
|
86 |
; |
|
87 |
dtspec: term '=' (con + '|') |
|
88 |
; |
|
89 |
con: name ('(' (term ',' +) ')')? |
|
90 |
; |
|
91 |
hints: monos? typeintros? typeelims? |
|
92 |
; |
|
93 |
\end{rail} |
|
94 |
||
95 |
See \cite{isabelle-ZF} for further information on inductive |
|
96 |
definitions in ZF, but note that this covers the old-style theory |
|
97 |
format. |
|
98 |
*} |
|
99 |
||
100 |
||
101 |
subsection {* Primitive recursive functions *} |
|
102 |
||
103 |
text {* |
|
104 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
105 |
@{command_def (ZF) "primrec"} & : & @{text "theory \<rightarrow> theory"} \\ |
26845 | 106 |
\end{matharray} |
107 |
||
108 |
\begin{rail} |
|
109 |
'primrec' (thmdecl? prop +) |
|
110 |
; |
|
111 |
\end{rail} |
|
112 |
*} |
|
113 |
||
114 |
||
115 |
subsection {* Cases and induction: emulating tactic scripts *} |
|
116 |
||
117 |
text {* |
|
118 |
The following important tactical tools of Isabelle/ZF have been |
|
119 |
ported to Isar. These should not be used in proper proof texts. |
|
120 |
||
121 |
\begin{matharray}{rcl} |
|
28761
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
122 |
@{method_def (ZF) case_tac}@{text "\<^sup>*"} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
123 |
@{method_def (ZF) induct_tac}@{text "\<^sup>*"} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
124 |
@{method_def (ZF) ind_cases}@{text "\<^sup>*"} & : & @{text method} \\ |
9ec4482c9201
updated/refined types of Isar language elements, removed special LaTeX macros;
wenzelm
parents:
28760
diff
changeset
|
125 |
@{command_def (ZF) "inductive_cases"} & : & @{text "theory \<rightarrow> theory"} \\ |
26845 | 126 |
\end{matharray} |
127 |
||
128 |
\begin{rail} |
|
129 |
('case\_tac' | 'induct\_tac') goalspec? name |
|
130 |
; |
|
131 |
indcases (prop +) |
|
132 |
; |
|
133 |
inductivecases (thmdecl? (prop +) + 'and') |
|
134 |
; |
|
135 |
\end{rail} |
|
136 |
*} |
|
137 |
||
26840 | 138 |
end |