author | wenzelm |
Thu, 10 Apr 2025 14:12:33 +0200 | |
changeset 82475 | 0a6d57c4d58b |
parent 76987 | 4c275405faae |
permissions | -rw-r--r-- |
48956 | 1 |
theory ZF_Isar |
65449
c82e63b11b8b
clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents:
62969
diff
changeset
|
2 |
imports ZF |
26840 | 3 |
begin |
4 |
||
48956 | 5 |
(*<*) |
69605 | 6 |
ML_file \<open>../antiquote_setup.ML\<close> |
48956 | 7 |
(*>*) |
8 |
||
67406 | 9 |
chapter \<open>Some Isar language elements\<close> |
26845 | 10 |
|
67406 | 11 |
section \<open>Type checking\<close> |
26845 | 12 |
|
67406 | 13 |
text \<open> |
26845 | 14 |
The ZF logic is essentially untyped, so the concept of ``type |
15 |
checking'' is performed as logical reasoning about set-membership |
|
16 |
statements. A special method assists users in this task; a version |
|
17 |
of this is already declared as a ``solver'' in the standard |
|
18 |
Simplifier setup. |
|
19 |
||
20 |
\begin{matharray}{rcl} |
|
69505 | 21 |
@{command_def (ZF) "print_tcset"}\<open>\<^sup>*\<close> & : & \<open>context \<rightarrow>\<close> \\ |
22 |
@{method_def (ZF) typecheck} & : & \<open>method\<close> \\ |
|
23 |
@{attribute_def (ZF) TC} & : & \<open>attribute\<close> \\ |
|
26845 | 24 |
\end{matharray} |
25 |
||
69593 | 26 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
27 |
@@{attribute (ZF) TC} (() | 'add' | 'del') |
69593 | 28 |
\<close> |
26845 | 29 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
30 |
\begin{description} |
26845 | 31 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
32 |
\item @{command (ZF) "print_tcset"} prints the collection of |
26845 | 33 |
typechecking rules of the current context. |
34 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
35 |
\item @{method (ZF) typecheck} attempts to solve any pending |
26845 | 36 |
type-checking problems in subgoals. |
37 |
||
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
38 |
\item @{attribute (ZF) TC} adds or deletes type-checking rules from |
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
39 |
the context. |
26845 | 40 |
|
28760
cbc435f7b16b
unified use of declaration environment with IsarImplementation;
wenzelm
parents:
26894
diff
changeset
|
41 |
\end{description} |
67406 | 42 |
\<close> |
26845 | 43 |
|
44 |
||
67406 | 45 |
section \<open>(Co)Inductive sets and datatypes\<close> |
26845 | 46 |
|
67406 | 47 |
subsection \<open>Set definitions\<close> |
26845 | 48 |
|
67406 | 49 |
text \<open> |
26845 | 50 |
In ZF everything is a set. The generic inductive package also |
51 |
provides a specific view for ``datatype'' specifications. |
|
52 |
Coinductive definitions are available in both cases, too. |
|
53 |
||
54 |
\begin{matharray}{rcl} |
|
69505 | 55 |
@{command_def (ZF) "inductive"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
56 |
@{command_def (ZF) "coinductive"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
57 |
@{command_def (ZF) "datatype"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
58 |
@{command_def (ZF) "codatatype"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
26845 | 59 |
\end{matharray} |
60 |
||
69593 | 61 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
62 |
(@@{command (ZF) inductive} | @@{command (ZF) coinductive}) domains intros hints |
26845 | 63 |
; |
64 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
65 |
domains: @'domains' (@{syntax term} + '+') ('<=' | '\<subseteq>') @{syntax term} |
26845 | 66 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
67 |
intros: @'intros' (@{syntax thmdecl}? @{syntax prop} +) |
26845 | 68 |
; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
69 |
hints: @{syntax (ZF) "monos"}? condefs? \<newline> |
42704 | 70 |
@{syntax (ZF) typeintros}? @{syntax (ZF) typeelims}? |
26845 | 71 |
; |
62969 | 72 |
@{syntax_def (ZF) "monos"}: @'monos' @{syntax thms} |
26845 | 73 |
; |
62969 | 74 |
condefs: @'con_defs' @{syntax thms} |
26845 | 75 |
; |
62969 | 76 |
@{syntax_def (ZF) typeintros}: @'type_intros' @{syntax thms} |
26845 | 77 |
; |
62969 | 78 |
@{syntax_def (ZF) typeelims}: @'type_elims' @{syntax thms} |
69593 | 79 |
\<close> |
26845 | 80 |
|
69505 | 81 |
In the following syntax specification \<open>monos\<close>, \<open>typeintros\<close>, and \<open>typeelims\<close> are the same as above. |
26845 | 82 |
|
69593 | 83 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
84 |
(@@{command (ZF) datatype} | @@{command (ZF) codatatype}) domain? (dtspec + @'and') hints |
26845 | 85 |
; |
86 |
||
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
87 |
domain: ('<=' | '\<subseteq>') @{syntax term} |
26845 | 88 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
89 |
dtspec: @{syntax term} '=' (con + '|') |
26845 | 90 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
91 |
con: @{syntax name} ('(' (@{syntax term} ',' +) ')')? |
26845 | 92 |
; |
55112
b1a5d603fd12
prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents:
55029
diff
changeset
|
93 |
hints: @{syntax (ZF) "monos"}? @{syntax (ZF) typeintros}? @{syntax (ZF) typeelims}? |
69593 | 94 |
\<close> |
26845 | 95 |
|
76987 | 96 |
See \<^cite>\<open>"isabelle-ZF"\<close> for further information on inductive |
26845 | 97 |
definitions in ZF, but note that this covers the old-style theory |
98 |
format. |
|
67406 | 99 |
\<close> |
26845 | 100 |
|
101 |
||
67406 | 102 |
subsection \<open>Primitive recursive functions\<close> |
26845 | 103 |
|
67406 | 104 |
text \<open> |
26845 | 105 |
\begin{matharray}{rcl} |
69505 | 106 |
@{command_def (ZF) "primrec"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
26845 | 107 |
\end{matharray} |
108 |
||
69593 | 109 |
\<^rail>\<open> |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
110 |
@@{command (ZF) primrec} (@{syntax thmdecl}? @{syntax prop} +) |
69593 | 111 |
\<close> |
67406 | 112 |
\<close> |
26845 | 113 |
|
114 |
||
67406 | 115 |
subsection \<open>Cases and induction: emulating tactic scripts\<close> |
26845 | 116 |
|
67406 | 117 |
text \<open> |
26845 | 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} |
|
69505 | 122 |
@{method_def (ZF) case_tac}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
123 |
@{method_def (ZF) induct_tac}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
|
124 |
@{method_def (ZF) ind_cases}\<open>\<^sup>*\<close> & : & \<open>method\<close> \\ |
|
125 |
@{command_def (ZF) "inductive_cases"} & : & \<open>theory \<rightarrow> theory\<close> \\ |
|
26845 | 126 |
\end{matharray} |
127 |
||
69593 | 128 |
\<^rail>\<open> |
42705 | 129 |
(@@{method (ZF) case_tac} | @@{method (ZF) induct_tac}) @{syntax goal_spec}? @{syntax name} |
26845 | 130 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
131 |
@@{method (ZF) ind_cases} (@{syntax prop} +) |
26845 | 132 |
; |
42596
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
wenzelm
parents:
40255
diff
changeset
|
133 |
@@{command (ZF) inductive_cases} (@{syntax thmdecl}? (@{syntax prop} +) + @'and') |
69593 | 134 |
\<close> |
67406 | 135 |
\<close> |
26845 | 136 |
|
26840 | 137 |
end |