| author | blanchet |
| Wed, 04 Nov 2015 15:07:23 +0100 | |
| changeset 61569 | 947ce60a06e1 |
| parent 61493 | 0debd22f0c0e |
| child 61656 | cfabbc083977 |
| permissions | -rw-r--r-- |
| 26779 | 1 |
theory Quick_Reference |
| 42651 | 2 |
imports Base Main |
| 26779 | 3 |
begin |
4 |
||
| 58618 | 5 |
chapter \<open>Isabelle/Isar quick reference \label{ap:refcard}\<close>
|
| 26779 | 6 |
|
| 58618 | 7 |
section \<open>Proof commands\<close> |
| 26779 | 8 |
|
| 58618 | 9 |
subsection \<open>Primitives and basic syntax\<close> |
| 26779 | 10 |
|
| 58618 | 11 |
text \<open> |
| 26779 | 12 |
\begin{tabular}{ll}
|
| 61493 | 13 |
@{command "fix"}~\<open>x\<close> & augment context by \<open>\<And>x. \<box>\<close> \\
|
14 |
@{command "assume"}~\<open>a: \<phi>\<close> & augment context by \<open>\<phi> \<Longrightarrow> \<box>\<close> \\
|
|
| 26779 | 15 |
@{command "then"} & indicate forward chaining of facts \\
|
| 61493 | 16 |
@{command "have"}~\<open>a: \<phi>\<close> & prove local result \\
|
17 |
@{command "show"}~\<open>a: \<phi>\<close> & prove local result, refining some goal \\
|
|
18 |
@{command "using"}~\<open>a\<close> & indicate use of additional facts \\
|
|
19 |
@{command "unfolding"}~\<open>a\<close> & unfold definitional equations \\
|
|
20 |
@{command "proof"}~\<open>m\<^sub>1\<close>~\dots~@{command "qed"}~\<open>m\<^sub>2\<close> & indicate proof structure and refinements \\
|
|
21 |
@{command "{"}~\<open>\<dots>\<close>~@{command "}"} & indicate explicit blocks \\
|
|
| 26779 | 22 |
@{command "next"} & switch blocks \\
|
| 61493 | 23 |
@{command "note"}~\<open>a = b\<close> & reconsider facts \\
|
24 |
@{command "let"}~\<open>p = t\<close> & abbreviate terms by higher-order matching \\
|
|
25 |
@{command "write"}~\<open>c (mx)\<close> & declare local mixfix syntax \\
|
|
| 26779 | 26 |
\end{tabular}
|
27 |
||
| 61421 | 28 |
\<^medskip> |
| 26852 | 29 |
|
30 |
\begin{tabular}{rcl}
|
|
| 61493 | 31 |
\<open>proof\<close> & = & \<open>prfx\<^sup>*\<close>~@{command "proof"}~\<open>method\<^sup>? stmt\<^sup>*\<close>~@{command "qed"}~\<open>method\<^sup>?\<close> \\
|
32 |
& \<open>|\<close> & \<open>prfx\<^sup>*\<close>~@{command "done"} \\
|
|
33 |
\<open>prfx\<close> & = & @{command "apply"}~\<open>method\<close> \\
|
|
34 |
& \<open>|\<close> & @{command "using"}~\<open>facts\<close> \\
|
|
35 |
& \<open>|\<close> & @{command "unfolding"}~\<open>facts\<close> \\
|
|
36 |
\<open>stmt\<close> & = & @{command "{"}~\<open>stmt\<^sup>*\<close>~@{command "}"} \\
|
|
37 |
& \<open>|\<close> & @{command "next"} \\
|
|
38 |
& \<open>|\<close> & @{command "note"}~\<open>name = facts\<close> \\
|
|
39 |
& \<open>|\<close> & @{command "let"}~\<open>term = term\<close> \\
|
|
40 |
& \<open>|\<close> & @{command "write"}~\<open>name (mixfix)\<close> \\
|
|
41 |
& \<open>|\<close> & @{command "fix"}~\<open>var\<^sup>+\<close> \\
|
|
42 |
& \<open>|\<close> & @{command "assume"}~\<open>name: props\<close> \\
|
|
43 |
& \<open>|\<close> & @{command "then"}\<open>\<^sup>?\<close>~\<open>goal\<close> \\
|
|
44 |
\<open>goal\<close> & = & @{command "have"}~\<open>name: props proof\<close> \\
|
|
45 |
& \<open>|\<close> & @{command "show"}~\<open>name: props proof\<close> \\
|
|
| 26852 | 46 |
\end{tabular}
|
| 58618 | 47 |
\<close> |
| 26779 | 48 |
|
49 |
||
| 58618 | 50 |
subsection \<open>Abbreviations and synonyms\<close> |
| 26779 | 51 |
|
| 58618 | 52 |
text \<open> |
| 26852 | 53 |
\begin{tabular}{rcl}
|
| 61493 | 54 |
@{command "by"}~\<open>m\<^sub>1 m\<^sub>2\<close> & \<open>\<equiv>\<close> &
|
55 |
@{command "proof"}~\<open>m\<^sub>1\<close>~@{command "qed"}~\<open>m\<^sub>2\<close> \\
|
|
56 |
@{command ".."} & \<open>\<equiv>\<close> & @{command "by"}~\<open>standard\<close> \\
|
|
57 |
@{command "."} & \<open>\<equiv>\<close> & @{command "by"}~\<open>this\<close> \\
|
|
58 |
@{command "hence"} & \<open>\<equiv>\<close> & @{command "then"}~@{command "have"} \\
|
|
59 |
@{command "thus"} & \<open>\<equiv>\<close> & @{command "then"}~@{command "show"} \\
|
|
60 |
@{command "from"}~\<open>a\<close> & \<open>\<equiv>\<close> & @{command "note"}~\<open>a\<close>~@{command "then"} \\
|
|
61 |
@{command "with"}~\<open>a\<close> & \<open>\<equiv>\<close> & @{command "from"}~\<open>a \<AND> this\<close> \\
|
|
62 |
@{command "from"}~\<open>this\<close> & \<open>\<equiv>\<close> & @{command "then"} \\
|
|
63 |
@{command "from"}~\<open>this\<close>~@{command "have"} & \<open>\<equiv>\<close> & @{command "hence"} \\
|
|
64 |
@{command "from"}~\<open>this\<close>~@{command "show"} & \<open>\<equiv>\<close> & @{command "thus"} \\
|
|
| 26852 | 65 |
\end{tabular}
|
| 58618 | 66 |
\<close> |
| 26779 | 67 |
|
68 |
||
| 58618 | 69 |
subsection \<open>Derived elements\<close> |
| 26779 | 70 |
|
| 58618 | 71 |
text \<open> |
| 26852 | 72 |
\begin{tabular}{rcl}
|
| 61493 | 73 |
@{command "also"}\<open>\<^sub>0\<close> & \<open>\<approx>\<close> &
|
74 |
@{command "note"}~\<open>calculation = this\<close> \\
|
|
75 |
@{command "also"}\<open>\<^sub>n\<^sub>+\<^sub>1\<close> & \<open>\<approx>\<close> &
|
|
76 |
@{command "note"}~\<open>calculation = trans [OF calculation this]\<close> \\
|
|
77 |
@{command "finally"} & \<open>\<approx>\<close> &
|
|
78 |
@{command "also"}~@{command "from"}~\<open>calculation\<close> \\[0.5ex]
|
|
79 |
@{command "moreover"} & \<open>\<approx>\<close> &
|
|
80 |
@{command "note"}~\<open>calculation = calculation this\<close> \\
|
|
81 |
@{command "ultimately"} & \<open>\<approx>\<close> &
|
|
82 |
@{command "moreover"}~@{command "from"}~\<open>calculation\<close> \\[0.5ex]
|
|
83 |
@{command "presume"}~\<open>a: \<phi>\<close> & \<open>\<approx>\<close> &
|
|
84 |
@{command "assume"}~\<open>a: \<phi>\<close> \\
|
|
85 |
@{command "def"}~\<open>a: x \<equiv> t\<close> & \<open>\<approx>\<close> &
|
|
86 |
@{command "fix"}~\<open>x\<close>~@{command "assume"}~\<open>a: x \<equiv> t\<close> \\
|
|
87 |
@{command "obtain"}~\<open>x \<WHERE> a: \<phi>\<close> & \<open>\<approx>\<close> &
|
|
88 |
\<open>\<dots>\<close>~@{command "fix"}~\<open>x\<close>~@{command "assume"}~\<open>a: \<phi>\<close> \\
|
|
89 |
@{command "case"}~\<open>c\<close> & \<open>\<approx>\<close> &
|
|
90 |
@{command "fix"}~\<open>x\<close>~@{command "assume"}~\<open>c: \<phi>\<close> \\
|
|
91 |
@{command "sorry"} & \<open>\<approx>\<close> &
|
|
92 |
@{command "by"}~\<open>cheating\<close> \\
|
|
| 26852 | 93 |
\end{tabular}
|
| 58618 | 94 |
\<close> |
| 26779 | 95 |
|
96 |
||
| 58618 | 97 |
subsection \<open>Diagnostic commands\<close> |
| 26779 | 98 |
|
| 58618 | 99 |
text \<open> |
| 26779 | 100 |
\begin{tabular}{ll}
|
| 57479 | 101 |
@{command "print_state"} & print proof state \\
|
102 |
@{command "print_statement"} & print fact in long statement form \\
|
|
| 61493 | 103 |
@{command "thm"}~\<open>a\<close> & print fact \\
|
104 |
@{command "prop"}~\<open>\<phi>\<close> & print proposition \\
|
|
105 |
@{command "term"}~\<open>t\<close> & print term \\
|
|
106 |
@{command "typ"}~\<open>\<tau>\<close> & print type \\
|
|
| 26779 | 107 |
\end{tabular}
|
| 58618 | 108 |
\<close> |
| 26779 | 109 |
|
110 |
||
| 58618 | 111 |
section \<open>Proof methods\<close> |
| 26779 | 112 |
|
| 58618 | 113 |
text \<open> |
| 26779 | 114 |
\begin{tabular}{ll}
|
| 61477 | 115 |
\multicolumn{2}{l}{\<^bold>\<open>Single steps (forward-chaining facts)\<close>} \\[0.5ex]
|
| 26779 | 116 |
@{method assumption} & apply some assumption \\
|
117 |
@{method this} & apply current facts \\
|
|
| 61493 | 118 |
@{method rule}~\<open>a\<close> & apply some rule \\
|
|
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
58618
diff
changeset
|
119 |
@{method standard} & apply standard rule (default for @{command "proof"}) \\
|
| 61493 | 120 |
@{method contradiction} & apply \<open>\<not>\<close> elimination rule (any order) \\
|
121 |
@{method cases}~\<open>t\<close> & case analysis (provides cases) \\
|
|
122 |
@{method induct}~\<open>x\<close> & proof by induction (provides cases) \\[2ex]
|
|
| 26779 | 123 |
|
| 61477 | 124 |
\multicolumn{2}{l}{\<^bold>\<open>Repeated steps (inserting facts)\<close>} \\[0.5ex]
|
| 26779 | 125 |
@{method "-"} & no rules \\
|
| 61493 | 126 |
@{method intro}~\<open>a\<close> & introduction rules \\
|
| 26779 | 127 |
@{method intro_classes} & class introduction rules \\
|
| 61493 | 128 |
@{method elim}~\<open>a\<close> & elimination rules \\
|
129 |
@{method unfold}~\<open>a\<close> & definitional rewrite rules \\[2ex]
|
|
| 26779 | 130 |
|
| 61477 | 131 |
\multicolumn{2}{l}{\<^bold>\<open>Automated proof tools (inserting facts)\<close>} \\[0.5ex]
|
| 26779 | 132 |
@{method iprover} & intuitionistic proof search \\
|
133 |
@{method blast}, @{method fast} & Classical Reasoner \\
|
|
134 |
@{method simp}, @{method simp_all} & Simplifier (+ Splitter) \\
|
|
135 |
@{method auto}, @{method force} & Simplifier + Classical Reasoner \\
|
|
136 |
@{method arith} & Arithmetic procedures \\
|
|
137 |
\end{tabular}
|
|
| 58618 | 138 |
\<close> |
| 26779 | 139 |
|
140 |
||
| 58618 | 141 |
section \<open>Attributes\<close> |
| 26779 | 142 |
|
| 58618 | 143 |
text \<open> |
| 26779 | 144 |
\begin{tabular}{ll}
|
| 61477 | 145 |
\multicolumn{2}{l}{\<^bold>\<open>Rules\<close>} \\[0.5ex]
|
| 61493 | 146 |
@{attribute OF}~\<open>a\<close> & rule resolved with facts (skipping ``\<open>_\<close>'') \\
|
147 |
@{attribute of}~\<open>t\<close> & rule instantiated with terms (skipping ``\<open>_\<close>'') \\
|
|
148 |
@{attribute "where"}~\<open>x = t\<close> & rule instantiated with terms, by variable name \\
|
|
| 26779 | 149 |
@{attribute symmetric} & resolution with symmetry rule \\
|
| 61493 | 150 |
@{attribute THEN}~\<open>b\<close> & resolution with another rule \\
|
| 26779 | 151 |
@{attribute rule_format} & result put into standard rule format \\
|
152 |
@{attribute elim_format} & destruct rule turned into elimination rule format \\[1ex]
|
|
153 |
||
| 61477 | 154 |
\multicolumn{2}{l}{\<^bold>\<open>Declarations\<close>} \\[0.5ex]
|
| 26779 | 155 |
@{attribute simp} & Simplifier rule \\
|
156 |
@{attribute intro}, @{attribute elim}, @{attribute dest} & Pure or Classical Reasoner rule \\
|
|
157 |
@{attribute iff} & Simplifier + Classical Reasoner rule \\
|
|
158 |
@{attribute split} & case split rule \\
|
|
159 |
@{attribute trans} & transitivity rule \\
|
|
160 |
@{attribute sym} & symmetry rule \\
|
|
161 |
\end{tabular}
|
|
| 58618 | 162 |
\<close> |
| 26779 | 163 |
|
164 |
||
| 58618 | 165 |
section \<open>Rule declarations and methods\<close> |
| 26779 | 166 |
|
| 58618 | 167 |
text \<open> |
| 26779 | 168 |
\begin{tabular}{l|lllll}
|
169 |
& @{method rule} & @{method iprover} & @{method blast} & @{method simp} & @{method auto} \\
|
|
170 |
& & & @{method fast} & @{method simp_all} & @{method force} \\
|
|
171 |
\hline |
|
| 61493 | 172 |
@{attribute Pure.elim}\<open>!\<close> @{attribute Pure.intro}\<open>!\<close>
|
173 |
& \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
|
| 26779 | 174 |
@{attribute Pure.elim} @{attribute Pure.intro}
|
| 61493 | 175 |
& \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
176 |
@{attribute elim}\<open>!\<close> @{attribute intro}\<open>!\<close>
|
|
177 |
& \<open>\<times>\<close> & & \<open>\<times>\<close> & & \<open>\<times>\<close> \\ |
|
| 26779 | 178 |
@{attribute elim} @{attribute intro}
|
| 61493 | 179 |
& \<open>\<times>\<close> & & \<open>\<times>\<close> & & \<open>\<times>\<close> \\ |
| 26779 | 180 |
@{attribute iff}
|
| 61493 | 181 |
& \<open>\<times>\<close> & & \<open>\<times>\<close> & \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
182 |
@{attribute iff}\<open>?\<close>
|
|
183 |
& \<open>\<times>\<close> \\ |
|
184 |
@{attribute elim}\<open>?\<close> @{attribute intro}\<open>?\<close>
|
|
185 |
& \<open>\<times>\<close> \\ |
|
| 26779 | 186 |
@{attribute simp}
|
| 61493 | 187 |
& & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
| 26779 | 188 |
@{attribute cong}
|
| 61493 | 189 |
& & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
| 26779 | 190 |
@{attribute split}
|
| 61493 | 191 |
& & & & \<open>\<times>\<close> & \<open>\<times>\<close> \\ |
| 26779 | 192 |
\end{tabular}
|
| 58618 | 193 |
\<close> |
| 26779 | 194 |
|
195 |
||
| 58618 | 196 |
section \<open>Emulating tactic scripts\<close> |
| 26779 | 197 |
|
| 58618 | 198 |
subsection \<open>Commands\<close> |
| 26779 | 199 |
|
| 58618 | 200 |
text \<open> |
| 26779 | 201 |
\begin{tabular}{ll}
|
| 61493 | 202 |
@{command "apply"}~\<open>m\<close> & apply proof method at initial position \\
|
203 |
@{command "apply_end"}~\<open>m\<close> & apply proof method near terminal position \\
|
|
| 26779 | 204 |
@{command "done"} & complete proof \\
|
| 61493 | 205 |
@{command "defer"}~\<open>n\<close> & move subgoal to end \\
|
206 |
@{command "prefer"}~\<open>n\<close> & move subgoal to beginning \\
|
|
| 26779 | 207 |
@{command "back"} & backtrack last command \\
|
208 |
\end{tabular}
|
|
| 58618 | 209 |
\<close> |
| 26779 | 210 |
|
211 |
||
| 58618 | 212 |
subsection \<open>Methods\<close> |
| 26779 | 213 |
|
| 58618 | 214 |
text \<open> |
| 26779 | 215 |
\begin{tabular}{ll}
|
| 61493 | 216 |
@{method rule_tac}~\<open>insts\<close> & resolution (with instantiation) \\
|
217 |
@{method erule_tac}~\<open>insts\<close> & elim-resolution (with instantiation) \\
|
|
218 |
@{method drule_tac}~\<open>insts\<close> & destruct-resolution (with instantiation) \\
|
|
219 |
@{method frule_tac}~\<open>insts\<close> & forward-resolution (with instantiation) \\
|
|
220 |
@{method cut_tac}~\<open>insts\<close> & insert facts (with instantiation) \\
|
|
221 |
@{method thin_tac}~\<open>\<phi>\<close> & delete assumptions \\
|
|
222 |
@{method subgoal_tac}~\<open>\<phi>\<close> & new claims \\
|
|
223 |
@{method rename_tac}~\<open>x\<close> & rename innermost goal parameters \\
|
|
224 |
@{method rotate_tac}~\<open>n\<close> & rotate assumptions of goal \\
|
|
225 |
@{method tactic}~\<open>text\<close> & arbitrary ML tactic \\
|
|
226 |
@{method case_tac}~\<open>t\<close> & exhaustion (datatypes) \\
|
|
227 |
@{method induct_tac}~\<open>x\<close> & induction (datatypes) \\
|
|
228 |
@{method ind_cases}~\<open>t\<close> & exhaustion + simplification (inductive predicates) \\
|
|
| 26779 | 229 |
\end{tabular}
|
| 58618 | 230 |
\<close> |
| 26779 | 231 |
|
232 |
end |