| author | huffman | 
| Sat, 03 Sep 2011 16:00:09 -0700 | |
| changeset 44695 | 075327b8e841 | 
| parent 42705 | 528a2ba8fa74 | 
| permissions | -rw-r--r-- | 
| 26840 | 1 | % | 
| 2 | \begin{isabellebody}%
 | |
| 40406 | 3 | \def\isabellecontext{ZF{\isaliteral{5F}{\isacharunderscore}}Specific}%
 | 
| 26840 | 4 | % | 
| 5 | \isadelimtheory | |
| 6 | % | |
| 7 | \endisadelimtheory | |
| 8 | % | |
| 9 | \isatagtheory | |
| 10 | \isacommand{theory}\isamarkupfalse%
 | |
| 40406 | 11 | \ ZF{\isaliteral{5F}{\isacharunderscore}}Specific\isanewline
 | 
| 42651 | 12 | \isakeyword{imports}\ Base\ Main\isanewline
 | 
| 26845 | 13 | \isakeyword{begin}%
 | 
| 14 | \endisatagtheory | |
| 15 | {\isafoldtheory}%
 | |
| 16 | % | |
| 17 | \isadelimtheory | |
| 18 | % | |
| 19 | \endisadelimtheory | |
| 20 | % | |
| 26852 | 21 | \isamarkupchapter{Isabelle/ZF \label{ch:zf}%
 | 
| 26845 | 22 | } | 
| 23 | \isamarkuptrue% | |
| 24 | % | |
| 25 | \isamarkupsection{Type checking%
 | |
| 26 | } | |
| 27 | \isamarkuptrue% | |
| 28 | % | |
| 29 | \begin{isamarkuptext}%
 | |
| 30 | The ZF logic is essentially untyped, so the concept of ``type | |
| 31 | checking'' is performed as logical reasoning about set-membership | |
| 32 | statements. A special method assists users in this task; a version | |
| 33 | of this is already declared as a ``solver'' in the standard | |
| 34 | Simplifier setup. | |
| 35 | ||
| 36 |   \begin{matharray}{rcl}
 | |
| 40406 | 37 |     \indexdef{ZF}{command}{print\_tcset}\hypertarget{command.ZF.print-tcset}{\hyperlink{command.ZF.print-tcset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}tcset}}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}context\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 28788 | 38 |     \indexdef{ZF}{method}{typecheck}\hypertarget{method.ZF.typecheck}{\hyperlink{method.ZF.typecheck}{\mbox{\isa{typecheck}}}} & : & \isa{method} \\
 | 
| 39 |     \indexdef{ZF}{attribute}{TC}\hypertarget{attribute.ZF.TC}{\hyperlink{attribute.ZF.TC}{\mbox{\isa{TC}}}} & : & \isa{attribute} \\
 | |
| 26845 | 40 |   \end{matharray}
 | 
| 41 | ||
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 42 |   \begin{railoutput}
 | 
| 42662 | 43 | \rail@begin{3}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 44 | \rail@term{\hyperlink{attribute.ZF.TC}{\mbox{\isa{TC}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 45 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 46 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 47 | \rail@term{\isa{add}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 48 | \rail@nextbar{2}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 49 | \rail@term{\isa{del}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 50 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 51 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 52 | \end{railoutput}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 53 | |
| 26845 | 54 | |
| 28788 | 55 |   \begin{description}
 | 
| 26845 | 56 | |
| 40406 | 57 |   \item \hyperlink{command.ZF.print-tcset}{\mbox{\isa{\isacommand{print{\isaliteral{5F}{\isacharunderscore}}tcset}}}} prints the collection of
 | 
| 26845 | 58 | typechecking rules of the current context. | 
| 59 | ||
| 28788 | 60 |   \item \hyperlink{method.ZF.typecheck}{\mbox{\isa{typecheck}}} attempts to solve any pending
 | 
| 26845 | 61 | type-checking problems in subgoals. | 
| 62 | ||
| 28788 | 63 |   \item \hyperlink{attribute.ZF.TC}{\mbox{\isa{TC}}} adds or deletes type-checking rules from
 | 
| 64 | the context. | |
| 26845 | 65 | |
| 28788 | 66 |   \end{description}%
 | 
| 26845 | 67 | \end{isamarkuptext}%
 | 
| 68 | \isamarkuptrue% | |
| 69 | % | |
| 70 | \isamarkupsection{(Co)Inductive sets and datatypes%
 | |
| 71 | } | |
| 72 | \isamarkuptrue% | |
| 73 | % | |
| 74 | \isamarkupsubsection{Set definitions%
 | |
| 75 | } | |
| 76 | \isamarkuptrue% | |
| 77 | % | |
| 78 | \begin{isamarkuptext}%
 | |
| 79 | In ZF everything is a set. The generic inductive package also | |
| 80 | provides a specific view for ``datatype'' specifications. | |
| 81 | Coinductive definitions are available in both cases, too. | |
| 82 | ||
| 83 |   \begin{matharray}{rcl}
 | |
| 40406 | 84 |     \indexdef{ZF}{command}{inductive}\hypertarget{command.ZF.inductive}{\hyperlink{command.ZF.inductive}{\mbox{\isa{\isacommand{inductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 85 |     \indexdef{ZF}{command}{coinductive}\hypertarget{command.ZF.coinductive}{\hyperlink{command.ZF.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | |
| 86 |     \indexdef{ZF}{command}{datatype}\hypertarget{command.ZF.datatype}{\hyperlink{command.ZF.datatype}{\mbox{\isa{\isacommand{datatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | |
| 87 |     \indexdef{ZF}{command}{codatatype}\hypertarget{command.ZF.codatatype}{\hyperlink{command.ZF.codatatype}{\mbox{\isa{\isacommand{codatatype}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | |
| 26845 | 88 |   \end{matharray}
 | 
| 89 | ||
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 90 |   \begin{railoutput}
 | 
| 42662 | 91 | \rail@begin{2}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 92 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 93 | \rail@term{\hyperlink{command.ZF.inductive}{\mbox{\isa{\isacommand{inductive}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 94 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 95 | \rail@term{\hyperlink{command.ZF.coinductive}{\mbox{\isa{\isacommand{coinductive}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 96 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 97 | \rail@nont{\isa{domains}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 98 | \rail@nont{\isa{intros}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 99 | \rail@nont{\isa{hints}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 100 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 101 | \rail@begin{2}{\isa{domains}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 102 | \rail@term{\isa{\isakeyword{domains}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 103 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 104 | \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 105 | \rail@nextplus{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 106 | \rail@cterm{\isa{{\isaliteral{2B}{\isacharplus}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 107 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 108 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 109 | \rail@term{\isa{{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 110 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 111 | \rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 112 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 113 | \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 114 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 115 | \rail@begin{3}{\isa{intros}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 116 | \rail@term{\isa{\isakeyword{intros}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 117 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 118 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 119 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 120 | \rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 121 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 122 | \rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 123 | \rail@nextplus{2}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 124 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 125 | \rail@end | 
| 42704 | 126 | \rail@begin{5}{\isa{hints}}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 127 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 128 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 129 | \rail@nont{\hyperlink{syntax.ZF.monos}{\mbox{\isa{monos}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 130 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 131 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 132 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 133 | \rail@nont{\isa{condefs}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 134 | \rail@endbar | 
| 42704 | 135 | \rail@cr{3}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 136 | \rail@bar | 
| 42704 | 137 | \rail@nextbar{4}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 138 | \rail@nont{\hyperlink{syntax.ZF.typeintros}{\mbox{\isa{typeintros}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 139 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 140 | \rail@bar | 
| 42704 | 141 | \rail@nextbar{4}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 142 | \rail@nont{\hyperlink{syntax.ZF.typeelims}{\mbox{\isa{typeelims}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 143 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 144 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 145 | \rail@begin{1}{\indexdef{ZF}{syntax}{monos}\hypertarget{syntax.ZF.monos}{\hyperlink{syntax.ZF.monos}{\mbox{\isa{monos}}}}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 146 | \rail@term{\isa{\isakeyword{monos}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 147 | \rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 148 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 149 | \rail@begin{1}{\isa{condefs}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 150 | \rail@term{\isa{\isakeyword{con{\isaliteral{5F}{\isacharunderscore}}defs}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 151 | \rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 152 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 153 | \rail@begin{1}{\indexdef{ZF}{syntax}{typeintros}\hypertarget{syntax.ZF.typeintros}{\hyperlink{syntax.ZF.typeintros}{\mbox{\isa{typeintros}}}}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 154 | \rail@term{\isa{\isakeyword{type{\isaliteral{5F}{\isacharunderscore}}intros}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 155 | \rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 156 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 157 | \rail@begin{1}{\indexdef{ZF}{syntax}{typeelims}\hypertarget{syntax.ZF.typeelims}{\hyperlink{syntax.ZF.typeelims}{\mbox{\isa{typeelims}}}}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 158 | \rail@term{\isa{\isakeyword{type{\isaliteral{5F}{\isacharunderscore}}elims}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 159 | \rail@nont{\hyperlink{syntax.thmrefs}{\mbox{\isa{thmrefs}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 160 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 161 | \end{railoutput}
 | 
| 26845 | 162 | |
| 163 | ||
| 40406 | 164 |   In the following syntax specification \isa{{\isaliteral{22}{\isachardoublequote}}monos{\isaliteral{22}{\isachardoublequote}}}, \isa{typeintros}, and \isa{typeelims} are the same as above.
 | 
| 26845 | 165 | |
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 166 |   \begin{railoutput}
 | 
| 42662 | 167 | \rail@begin{2}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 168 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 169 | \rail@term{\hyperlink{command.ZF.datatype}{\mbox{\isa{\isacommand{datatype}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 170 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 171 | \rail@term{\hyperlink{command.ZF.codatatype}{\mbox{\isa{\isacommand{codatatype}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 172 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 173 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 174 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 175 | \rail@nont{\isa{domain}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 176 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 177 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 178 | \rail@nont{\isa{dtspec}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 179 | \rail@nextplus{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 180 | \rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 181 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 182 | \rail@nont{\isa{hints}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 183 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 184 | \rail@begin{2}{\isa{domain}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 185 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 186 | \rail@term{\isa{{\isaliteral{3C}{\isacharless}}{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 187 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 188 | \rail@term{\isa{{\isaliteral{5C3C73756273657465713E}{\isasymsubseteq}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 189 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 190 | \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 191 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 192 | \rail@begin{2}{\isa{dtspec}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 193 | \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 194 | \rail@term{\isa{{\isaliteral{3D}{\isacharequal}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 195 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 196 | \rail@nont{\isa{con}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 197 | \rail@nextplus{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 198 | \rail@cterm{\isa{{\isaliteral{7C}{\isacharbar}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 199 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 200 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 201 | \rail@begin{3}{\isa{con}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 202 | \rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 203 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 204 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 205 | \rail@term{\isa{{\isaliteral{28}{\isacharparenleft}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 206 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 207 | \rail@nont{\hyperlink{syntax.term}{\mbox{\isa{term}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 208 | \rail@term{\isa{{\isaliteral{2C}{\isacharcomma}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 209 | \rail@nextplus{2}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 210 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 211 | \rail@term{\isa{{\isaliteral{29}{\isacharparenright}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 212 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 213 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 214 | \rail@begin{2}{\isa{hints}}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 215 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 216 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 217 | \rail@nont{\hyperlink{syntax.ZF.monos}{\mbox{\isa{monos}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 218 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 219 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 220 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 221 | \rail@nont{\hyperlink{syntax.ZF.typeintros}{\mbox{\isa{typeintros}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 222 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 223 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 224 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 225 | \rail@nont{\hyperlink{syntax.ZF.typeelims}{\mbox{\isa{typeelims}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 226 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 227 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 228 | \end{railoutput}
 | 
| 26845 | 229 | |
| 230 | ||
| 231 |   See \cite{isabelle-ZF} for further information on inductive
 | |
| 232 | definitions in ZF, but note that this covers the old-style theory | |
| 233 | format.% | |
| 234 | \end{isamarkuptext}%
 | |
| 235 | \isamarkuptrue% | |
| 236 | % | |
| 237 | \isamarkupsubsection{Primitive recursive functions%
 | |
| 238 | } | |
| 239 | \isamarkuptrue% | |
| 240 | % | |
| 241 | \begin{isamarkuptext}%
 | |
| 242 | \begin{matharray}{rcl}
 | |
| 40406 | 243 |     \indexdef{ZF}{command}{primrec}\hypertarget{command.ZF.primrec}{\hyperlink{command.ZF.primrec}{\mbox{\isa{\isacommand{primrec}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | 
| 26845 | 244 |   \end{matharray}
 | 
| 245 | ||
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 246 |   \begin{railoutput}
 | 
| 42662 | 247 | \rail@begin{3}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 248 | \rail@term{\hyperlink{command.ZF.primrec}{\mbox{\isa{\isacommand{primrec}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 249 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 250 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 251 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 252 | \rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 253 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 254 | \rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 255 | \rail@nextplus{2}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 256 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 257 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 258 | \end{railoutput}%
 | 
| 26845 | 259 | \end{isamarkuptext}%
 | 
| 260 | \isamarkuptrue% | |
| 261 | % | |
| 262 | \isamarkupsubsection{Cases and induction: emulating tactic scripts%
 | |
| 263 | } | |
| 264 | \isamarkuptrue% | |
| 265 | % | |
| 266 | \begin{isamarkuptext}%
 | |
| 267 | The following important tactical tools of Isabelle/ZF have been | |
| 268 | ported to Isar. These should not be used in proper proof texts. | |
| 269 | ||
| 270 |   \begin{matharray}{rcl}
 | |
| 40406 | 271 |     \indexdef{ZF}{method}{case\_tac}\hypertarget{method.ZF.case-tac}{\hyperlink{method.ZF.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | 
| 272 |     \indexdef{ZF}{method}{induct\_tac}\hypertarget{method.ZF.induct-tac}{\hyperlink{method.ZF.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | |
| 273 |     \indexdef{ZF}{method}{ind\_cases}\hypertarget{method.ZF.ind-cases}{\hyperlink{method.ZF.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}\isa{{\isaliteral{22}{\isachardoublequote}}\isaliteral{5C3C5E7375703E}{}\isactrlsup {\isaliteral{2A}{\isacharasterisk}}{\isaliteral{22}{\isachardoublequote}}} & : & \isa{method} \\
 | |
| 274 |     \indexdef{ZF}{command}{inductive\_cases}\hypertarget{command.ZF.inductive-cases}{\hyperlink{command.ZF.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}} & : & \isa{{\isaliteral{22}{\isachardoublequote}}theory\ {\isaliteral{5C3C72696768746172726F773E}{\isasymrightarrow}}\ theory{\isaliteral{22}{\isachardoublequote}}} \\
 | |
| 26845 | 275 |   \end{matharray}
 | 
| 276 | ||
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 277 |   \begin{railoutput}
 | 
| 42662 | 278 | \rail@begin{2}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 279 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 280 | \rail@term{\hyperlink{method.ZF.case-tac}{\mbox{\isa{case{\isaliteral{5F}{\isacharunderscore}}tac}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 281 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 282 | \rail@term{\hyperlink{method.ZF.induct-tac}{\mbox{\isa{induct{\isaliteral{5F}{\isacharunderscore}}tac}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 283 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 284 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 285 | \rail@nextbar{1}
 | 
| 42705 | 286 | \rail@nont{\hyperlink{syntax.goal-spec}{\mbox{\isa{goal{\isaliteral{5F}{\isacharunderscore}}spec}}}}[]
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 287 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 288 | \rail@nont{\hyperlink{syntax.name}{\mbox{\isa{name}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 289 | \rail@end | 
| 42662 | 290 | \rail@begin{2}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 291 | \rail@term{\hyperlink{method.ZF.ind-cases}{\mbox{\isa{ind{\isaliteral{5F}{\isacharunderscore}}cases}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 292 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 293 | \rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 294 | \rail@nextplus{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 295 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 296 | \rail@end | 
| 42662 | 297 | \rail@begin{3}{}
 | 
| 42596 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 298 | \rail@term{\hyperlink{command.ZF.inductive-cases}{\mbox{\isa{\isacommand{inductive{\isaliteral{5F}{\isacharunderscore}}cases}}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 299 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 300 | \rail@bar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 301 | \rail@nextbar{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 302 | \rail@nont{\hyperlink{syntax.thmdecl}{\mbox{\isa{thmdecl}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 303 | \rail@endbar | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 304 | \rail@plus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 305 | \rail@nont{\hyperlink{syntax.prop}{\mbox{\isa{prop}}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 306 | \rail@nextplus{1}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 307 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 308 | \rail@nextplus{2}
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 309 | \rail@cterm{\isa{\isakeyword{and}}}[]
 | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 310 | \rail@endplus | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 311 | \rail@end | 
| 
6c621a9d612a
modernized rail diagrams using @{rail} antiquotation;
 wenzelm parents: 
40406diff
changeset | 312 | \end{railoutput}%
 | 
| 26845 | 313 | \end{isamarkuptext}%
 | 
| 314 | \isamarkuptrue% | |
| 315 | % | |
| 316 | \isadelimtheory | |
| 317 | % | |
| 318 | \endisadelimtheory | |
| 319 | % | |
| 320 | \isatagtheory | |
| 26840 | 321 | \isacommand{end}\isamarkupfalse%
 | 
| 322 | % | |
| 323 | \endisatagtheory | |
| 324 | {\isafoldtheory}%
 | |
| 325 | % | |
| 326 | \isadelimtheory | |
| 327 | % | |
| 328 | \endisadelimtheory | |
| 26845 | 329 | \isanewline | 
| 26840 | 330 | \end{isabellebody}%
 | 
| 331 | %%% Local Variables: | |
| 332 | %%% mode: latex | |
| 333 | %%% TeX-master: "root" | |
| 334 | %%% End: |