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