author | wenzelm |
Sun, 18 Oct 2015 22:57:09 +0200 | |
changeset 61477 | e467ae7aa808 |
parent 61421 | e0825405d398 |
child 61493 | 0debd22f0c0e |
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} |
13 |
@{command "fix"}~@{text x} & augment context by @{text "\<And>x. \<box>"} \\ |
|
14 |
@{command "assume"}~@{text "a: \<phi>"} & augment context by @{text "\<phi> \<Longrightarrow> \<box>"} \\ |
|
15 |
@{command "then"} & indicate forward chaining of facts \\ |
|
16 |
@{command "have"}~@{text "a: \<phi>"} & prove local result \\ |
|
17 |
@{command "show"}~@{text "a: \<phi>"} & prove local result, refining some goal \\ |
|
18 |
@{command "using"}~@{text a} & indicate use of additional facts \\ |
|
19 |
@{command "unfolding"}~@{text a} & unfold definitional equations \\ |
|
20 |
@{command "proof"}~@{text "m\<^sub>1"}~\dots~@{command "qed"}~@{text "m\<^sub>2"} & indicate proof structure and refinements \\ |
|
26852 | 21 |
@{command "{"}~@{text "\<dots>"}~@{command "}"} & indicate explicit blocks \\ |
26779 | 22 |
@{command "next"} & switch blocks \\ |
23 |
@{command "note"}~@{text "a = b"} & reconsider facts \\ |
|
24 |
@{command "let"}~@{text "p = t"} & abbreviate terms by higher-order matching \\ |
|
40535 | 25 |
@{command "write"}~@{text "c (mx)"} & declare local mixfix syntax \\ |
26779 | 26 |
\end{tabular} |
27 |
||
61421 | 28 |
\<^medskip> |
26852 | 29 |
|
30 |
\begin{tabular}{rcl} |
|
29724 | 31 |
@{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\ |
42667 | 32 |
& @{text "|"} & @{text "prfx\<^sup>*"}~@{command "done"} \\ |
26779 | 33 |
@{text prfx} & = & @{command "apply"}~@{text method} \\ |
26852 | 34 |
& @{text "|"} & @{command "using"}~@{text "facts"} \\ |
35 |
& @{text "|"} & @{command "unfolding"}~@{text "facts"} \\ |
|
26779 | 36 |
@{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\ |
26852 | 37 |
& @{text "|"} & @{command "next"} \\ |
38 |
& @{text "|"} & @{command "note"}~@{text "name = facts"} \\ |
|
39 |
& @{text "|"} & @{command "let"}~@{text "term = term"} \\ |
|
40535 | 40 |
& @{text "|"} & @{command "write"}~@{text "name (mixfix)"} \\ |
26852 | 41 |
& @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\ |
42 |
& @{text "|"} & @{command "assume"}~@{text "name: props"} \\ |
|
43 |
& @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\ |
|
26780 | 44 |
@{text goal} & = & @{command "have"}~@{text "name: props proof"} \\ |
26852 | 45 |
& @{text "|"} & @{command "show"}~@{text "name: props proof"} \\ |
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} |
54 |
@{command "by"}~@{text "m\<^sub>1 m\<^sub>2"} & @{text "\<equiv>"} & |
|
55 |
@{command "proof"}~@{text "m\<^sub>1"}~@{command "qed"}~@{text "m\<^sub>2"} \\ |
|
60618
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
wenzelm
parents:
58618
diff
changeset
|
56 |
@{command ".."} & @{text "\<equiv>"} & @{command "by"}~@{text standard} \\ |
26852 | 57 |
@{command "."} & @{text "\<equiv>"} & @{command "by"}~@{text this} \\ |
58 |
@{command "hence"} & @{text "\<equiv>"} & @{command "then"}~@{command "have"} \\ |
|
59 |
@{command "thus"} & @{text "\<equiv>"} & @{command "then"}~@{command "show"} \\ |
|
60 |
@{command "from"}~@{text a} & @{text "\<equiv>"} & @{command "note"}~@{text a}~@{command "then"} \\ |
|
42667 | 61 |
@{command "with"}~@{text a} & @{text "\<equiv>"} & @{command "from"}~@{text "a \<AND> this"} \\ |
26852 | 62 |
@{command "from"}~@{text this} & @{text "\<equiv>"} & @{command "then"} \\ |
63 |
@{command "from"}~@{text this}~@{command "have"} & @{text "\<equiv>"} & @{command "hence"} \\ |
|
64 |
@{command "from"}~@{text this}~@{command "show"} & @{text "\<equiv>"} & @{command "thus"} \\ |
|
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} |
73 |
@{command "also"}@{text "\<^sub>0"} & @{text "\<approx>"} & |
|
74 |
@{command "note"}~@{text "calculation = this"} \\ |
|
75 |
@{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & @{text "\<approx>"} & |
|
76 |
@{command "note"}~@{text "calculation = trans [OF calculation this]"} \\ |
|
77 |
@{command "finally"} & @{text "\<approx>"} & |
|
78 |
@{command "also"}~@{command "from"}~@{text calculation} \\[0.5ex] |
|
79 |
@{command "moreover"} & @{text "\<approx>"} & |
|
80 |
@{command "note"}~@{text "calculation = calculation this"} \\ |
|
81 |
@{command "ultimately"} & @{text "\<approx>"} & |
|
82 |
@{command "moreover"}~@{command "from"}~@{text calculation} \\[0.5ex] |
|
83 |
@{command "presume"}~@{text "a: \<phi>"} & @{text "\<approx>"} & |
|
84 |
@{command "assume"}~@{text "a: \<phi>"} \\ |
|
85 |
@{command "def"}~@{text "a: x \<equiv> t"} & @{text "\<approx>"} & |
|
86 |
@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: x \<equiv> t"} \\ |
|
87 |
@{command "obtain"}~@{text "x \<WHERE> a: \<phi>"} & @{text "\<approx>"} & |
|
88 |
@{text "\<dots>"}~@{command "fix"}~@{text x}~@{command "assume"}~@{text "a: \<phi>"} \\ |
|
89 |
@{command "case"}~@{text c} & @{text "\<approx>"} & |
|
90 |
@{command "fix"}~@{text x}~@{command "assume"}~@{text "c: \<phi>"} \\ |
|
91 |
@{command "sorry"} & @{text "\<approx>"} & |
|
92 |
@{command "by"}~@{text cheating} \\ |
|
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 \\ |
|
26779 | 103 |
@{command "thm"}~@{text a} & print fact \\ |
40536 | 104 |
@{command "prop"}~@{text \<phi>} & print proposition \\ |
26779 | 105 |
@{command "term"}~@{text t} & print term \\ |
40536 | 106 |
@{command "typ"}~@{text \<tau>} & 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 \\ |
|
118 |
@{method rule}~@{text a} & 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"}) \\ |
26779 | 120 |
@{method contradiction} & apply @{text "\<not>"} elimination rule (any order) \\ |
121 |
@{method cases}~@{text t} & case analysis (provides cases) \\ |
|
122 |
@{method induct}~@{text x} & proof by induction (provides cases) \\[2ex] |
|
123 |
||
61477 | 124 |
\multicolumn{2}{l}{\<^bold>\<open>Repeated steps (inserting facts)\<close>} \\[0.5ex] |
26779 | 125 |
@{method "-"} & no rules \\ |
126 |
@{method intro}~@{text a} & introduction rules \\ |
|
127 |
@{method intro_classes} & class introduction rules \\ |
|
128 |
@{method elim}~@{text a} & elimination rules \\ |
|
129 |
@{method unfold}~@{text a} & definitional rewrite rules \\[2ex] |
|
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] |
26779 | 146 |
@{attribute OF}~@{text a} & rule resolved with facts (skipping ``@{text _}'') \\ |
147 |
@{attribute of}~@{text t} & rule instantiated with terms (skipping ``@{text _}'') \\ |
|
148 |
@{attribute "where"}~@{text "x = t"} & rule instantiated with terms, by variable name \\ |
|
149 |
@{attribute symmetric} & resolution with symmetry rule \\ |
|
150 |
@{attribute THEN}~@{text b} & resolution with another rule \\ |
|
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 |
|
172 |
@{attribute Pure.elim}@{text "!"} @{attribute Pure.intro}@{text "!"} |
|
173 |
& @{text "\<times>"} & @{text "\<times>"} \\ |
|
174 |
@{attribute Pure.elim} @{attribute Pure.intro} |
|
175 |
& @{text "\<times>"} & @{text "\<times>"} \\ |
|
176 |
@{attribute elim}@{text "!"} @{attribute intro}@{text "!"} |
|
177 |
& @{text "\<times>"} & & @{text "\<times>"} & & @{text "\<times>"} \\ |
|
178 |
@{attribute elim} @{attribute intro} |
|
179 |
& @{text "\<times>"} & & @{text "\<times>"} & & @{text "\<times>"} \\ |
|
180 |
@{attribute iff} |
|
181 |
& @{text "\<times>"} & & @{text "\<times>"} & @{text "\<times>"} & @{text "\<times>"} \\ |
|
182 |
@{attribute iff}@{text "?"} |
|
183 |
& @{text "\<times>"} \\ |
|
184 |
@{attribute elim}@{text "?"} @{attribute intro}@{text "?"} |
|
185 |
& @{text "\<times>"} \\ |
|
186 |
@{attribute simp} |
|
187 |
& & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
188 |
@{attribute cong} |
|
189 |
& & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
190 |
@{attribute split} |
|
191 |
& & & & @{text "\<times>"} & @{text "\<times>"} \\ |
|
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} |
202 |
@{command "apply"}~@{text m} & apply proof method at initial position \\ |
|
203 |
@{command "apply_end"}~@{text m} & apply proof method near terminal position \\ |
|
204 |
@{command "done"} & complete proof \\ |
|
205 |
@{command "defer"}~@{text n} & move subgoal to end \\ |
|
206 |
@{command "prefer"}~@{text n} & move subgoal to beginning \\ |
|
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} |
216 |
@{method rule_tac}~@{text insts} & resolution (with instantiation) \\ |
|
217 |
@{method erule_tac}~@{text insts} & elim-resolution (with instantiation) \\ |
|
218 |
@{method drule_tac}~@{text insts} & destruct-resolution (with instantiation) \\ |
|
219 |
@{method frule_tac}~@{text insts} & forward-resolution (with instantiation) \\ |
|
220 |
@{method cut_tac}~@{text insts} & insert facts (with instantiation) \\ |
|
221 |
@{method thin_tac}~@{text \<phi>} & delete assumptions \\ |
|
222 |
@{method subgoal_tac}~@{text \<phi>} & new claims \\ |
|
223 |
@{method rename_tac}~@{text x} & rename innermost goal parameters \\ |
|
224 |
@{method rotate_tac}~@{text n} & rotate assumptions of goal \\ |
|
225 |
@{method tactic}~@{text "text"} & arbitrary ML tactic \\ |
|
226 |
@{method case_tac}~@{text t} & exhaustion (datatypes) \\ |
|
227 |
@{method induct_tac}~@{text x} & induction (datatypes) \\ |
|
228 |
@{method ind_cases}~@{text t} & exhaustion + simplification (inductive predicates) \\ |
|
229 |
\end{tabular} |
|
58618 | 230 |
\<close> |
26779 | 231 |
|
232 |
end |